moved Integ files to canonical place;
authorwenzelm
Thu, 31 May 2007 18:16:52 +0200
changeset 23164 69e55066dbca
parent 23163 eef345eff987
child 23165 5d319b0f8bf9
moved Integ files to canonical place;
src/HOL/IntArith.thy
src/HOL/IntDef.thy
src/HOL/IntDiv.thy
src/HOL/Integ/IntArith.thy
src/HOL/Integ/IntDef.thy
src/HOL/Integ/IntDiv.thy
src/HOL/Integ/NatBin.thy
src/HOL/Integ/NatSimprocs.thy
src/HOL/Integ/Numeral.thy
src/HOL/Integ/int_arith1.ML
src/HOL/Integ/int_factor_simprocs.ML
src/HOL/Integ/nat_simprocs.ML
src/HOL/IsaMakefile
src/HOL/NatBin.thy
src/HOL/NatSimprocs.thy
src/HOL/Numeral.thy
src/HOL/Presburger.thy
src/HOL/int_arith1.ML
src/HOL/int_factor_simprocs.ML
src/HOL/nat_simprocs.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IntArith.thy	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,411 @@
+(*  Title:      HOL/IntArith.thy
+    ID:         $Id$
+    Authors:    Larry Paulson and Tobias Nipkow
+*)
+
+header {* Integer arithmetic *}
+
+theory IntArith
+imports Numeral Wellfounded_Relations
+uses "~~/src/Provers/Arith/assoc_fold.ML" ("int_arith1.ML")
+begin
+
+text{*Duplicate: can't understand why it's necessary*}
+declare numeral_0_eq_0 [simp]
+
+
+subsection{*Inequality Reasoning for the Arithmetic Simproc*}
+
+lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)"
+by simp 
+
+lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)"
+by simp
+
+lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)"
+by simp 
+
+lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)"
+by simp
+
+lemma divide_numeral_1: "a / Numeral1 = (a::'a::{number_ring,field})"
+by simp
+
+lemma inverse_numeral_1:
+  "inverse Numeral1 = (Numeral1::'a::{number_ring,field})"
+by simp
+
+text{*Theorem lists for the cancellation simprocs. The use of binary numerals
+for 0 and 1 reduces the number of special cases.*}
+
+lemmas add_0s = add_numeral_0 add_numeral_0_right
+lemmas mult_1s = mult_numeral_1 mult_numeral_1_right 
+                 mult_minus1 mult_minus1_right
+
+
+subsection{*Special Arithmetic Rules for Abstract 0 and 1*}
+
+text{*Arithmetic computations are defined for binary literals, which leaves 0
+and 1 as special cases. Addition already has rules for 0, but not 1.
+Multiplication and unary minus already have rules for both 0 and 1.*}
+
+
+lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'"
+by simp
+
+
+lemmas add_number_of_eq = number_of_add [symmetric]
+
+text{*Allow 1 on either or both sides*}
+lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)"
+by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq)
+
+lemmas add_special =
+    one_add_one_is_two
+    binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard]
+
+text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*}
+lemmas diff_special =
+    binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard]
+
+text{*Allow 0 or 1 on either side with a binary numeral on the other*}
+lemmas eq_special =
+    binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard]
+    binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard]
+    binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard]
+
+text{*Allow 0 or 1 on either side with a binary numeral on the other*}
+lemmas less_special =
+  binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard]
+  binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard]
+  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard]
+  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard]
+
+text{*Allow 0 or 1 on either side with a binary numeral on the other*}
+lemmas le_special =
+    binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard]
+    binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard]
+    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard]
+    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard]
+
+lemmas arith_special[simp] = 
+       add_special diff_special eq_special less_special le_special
+
+
+lemma min_max_01: "min (0::int) 1 = 0 & min (1::int) 0 = 0 &
+                   max (0::int) 1 = 1 & max (1::int) 0 = 1"
+by(simp add:min_def max_def)
+
+lemmas min_max_special[simp] =
+ min_max_01
+ max_def[of "0::int" "number_of v", standard, simp]
+ min_def[of "0::int" "number_of v", standard, simp]
+ max_def[of "number_of u" "0::int", standard, simp]
+ min_def[of "number_of u" "0::int", standard, simp]
+ max_def[of "1::int" "number_of v", standard, simp]
+ min_def[of "1::int" "number_of v", standard, simp]
+ max_def[of "number_of u" "1::int", standard, simp]
+ min_def[of "number_of u" "1::int", standard, simp]
+
+use "int_arith1.ML"
+setup int_arith_setup
+
+
+subsection{*Lemmas About Small Numerals*}
+
+lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)"
+proof -
+  have "(of_int -1 :: 'a) = of_int (- 1)" by simp
+  also have "... = - of_int 1" by (simp only: of_int_minus)
+  also have "... = -1" by simp
+  finally show ?thesis .
+qed
+
+lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_idom,number_ring})"
+by (simp add: abs_if)
+
+lemma abs_power_minus_one [simp]:
+     "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,recpower})"
+by (simp add: power_abs)
+
+lemma of_int_number_of_eq:
+     "of_int (number_of v) = (number_of v :: 'a :: number_ring)"
+by (simp add: number_of_eq) 
+
+text{*Lemmas for specialist use, NOT as default simprules*}
+lemma mult_2: "2 * z = (z+z::'a::number_ring)"
+proof -
+  have "2*z = (1 + 1)*z" by simp
+  also have "... = z+z" by (simp add: left_distrib)
+  finally show ?thesis .
+qed
+
+lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)"
+by (subst mult_commute, rule mult_2)
+
+
+subsection{*More Inequality Reasoning*}
+
+lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)"
+by arith
+
+lemma add1_zle_eq: "(w + (1::int) \<le> z) = (w<z)"
+by arith
+
+lemma zle_diff1_eq [simp]: "(w \<le> z - (1::int)) = (w<z)"
+by arith
+
+lemma zle_add1_eq_le [simp]: "(w < z + (1::int)) = (w\<le>z)"
+by arith
+
+lemma int_one_le_iff_zero_less: "((1::int) \<le> z) = (0 < z)"
+by arith
+
+
+subsection{*The Functions @{term nat} and @{term int}*}
+
+text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and
+  @{term "w + - z"}*}
+declare Zero_int_def [symmetric, simp]
+declare One_int_def [symmetric, simp]
+
+lemmas diff_int_def_symmetric = diff_int_def [symmetric, simp]
+
+lemma nat_0: "nat 0 = 0"
+by (simp add: nat_eq_iff)
+
+lemma nat_1: "nat 1 = Suc 0"
+by (subst nat_eq_iff, simp)
+
+lemma nat_2: "nat 2 = Suc (Suc 0)"
+by (subst nat_eq_iff, simp)
+
+lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
+apply (insert zless_nat_conj [of 1 z])
+apply (auto simp add: nat_1)
+done
+
+text{*This simplifies expressions of the form @{term "int n = z"} where
+      z is an integer literal.*}
+lemmas int_eq_iff_number_of [simp] = int_eq_iff [of _ "number_of v", standard]
+
+
+lemma split_nat [arith_split]:
+  "P(nat(i::int)) = ((\<forall>n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))"
+  (is "?P = (?L & ?R)")
+proof (cases "i < 0")
+  case True thus ?thesis by simp
+next
+  case False
+  have "?P = ?L"
+  proof
+    assume ?P thus ?L using False by clarsimp
+  next
+    assume ?L thus ?P using False by simp
+  qed
+  with False show ?thesis by simp
+qed
+
+
+(*Analogous to zadd_int*)
+lemma zdiff_int: "n \<le> m ==> int m - int n = int (m-n)" 
+by (induct m n rule: diff_induct, simp_all)
+
+lemma nat_mult_distrib: "(0::int) \<le> z ==> nat (z*z') = nat z * nat z'"
+apply (cases "0 \<le> z'")
+apply (rule inj_int [THEN injD])
+apply (simp add: int_mult zero_le_mult_iff)
+apply (simp add: mult_le_0_iff)
+done
+
+lemma nat_mult_distrib_neg: "z \<le> (0::int) ==> nat(z*z') = nat(-z) * nat(-z')"
+apply (rule trans)
+apply (rule_tac [2] nat_mult_distrib, auto)
+done
+
+lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)"
+apply (cases "z=0 | w=0")
+apply (auto simp add: abs_if nat_mult_distrib [symmetric] 
+                      nat_mult_distrib_neg [symmetric] mult_less_0_iff)
+done
+
+
+subsection "Induction principles for int"
+
+text{*Well-founded segments of the integers*}
+
+definition
+  int_ge_less_than  ::  "int => (int * int) set"
+where
+  "int_ge_less_than d = {(z',z). d \<le> z' & z' < z}"
+
+theorem wf_int_ge_less_than: "wf (int_ge_less_than d)"
+proof -
+  have "int_ge_less_than d \<subseteq> measure (%z. nat (z-d))"
+    by (auto simp add: int_ge_less_than_def)
+  thus ?thesis 
+    by (rule wf_subset [OF wf_measure]) 
+qed
+
+text{*This variant looks odd, but is typical of the relations suggested
+by RankFinder.*}
+
+definition
+  int_ge_less_than2 ::  "int => (int * int) set"
+where
+  "int_ge_less_than2 d = {(z',z). d \<le> z & z' < z}"
+
+theorem wf_int_ge_less_than2: "wf (int_ge_less_than2 d)"
+proof -
+  have "int_ge_less_than2 d \<subseteq> measure (%z. nat (1+z-d))" 
+    by (auto simp add: int_ge_less_than2_def)
+  thus ?thesis 
+    by (rule wf_subset [OF wf_measure]) 
+qed
+
+                     (* `set:int': dummy construction *)
+theorem int_ge_induct[case_names base step,induct set:int]:
+  assumes ge: "k \<le> (i::int)" and
+        base: "P(k)" and
+        step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
+  shows "P i"
+proof -
+  { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i"
+    proof (induct n)
+      case 0
+      hence "i = k" by arith
+      thus "P i" using base by simp
+    next
+      case (Suc n)
+      hence "n = nat((i - 1) - k)" by arith
+      moreover
+      have ki1: "k \<le> i - 1" using Suc.prems by arith
+      ultimately
+      have "P(i - 1)" by(rule Suc.hyps)
+      from step[OF ki1 this] show ?case by simp
+    qed
+  }
+  with ge show ?thesis by fast
+qed
+
+                     (* `set:int': dummy construction *)
+theorem int_gr_induct[case_names base step,induct set:int]:
+  assumes gr: "k < (i::int)" and
+        base: "P(k+1)" and
+        step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
+  shows "P i"
+apply(rule int_ge_induct[of "k + 1"])
+  using gr apply arith
+ apply(rule base)
+apply (rule step, simp+)
+done
+
+theorem int_le_induct[consumes 1,case_names base step]:
+  assumes le: "i \<le> (k::int)" and
+        base: "P(k)" and
+        step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
+  shows "P i"
+proof -
+  { fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i"
+    proof (induct n)
+      case 0
+      hence "i = k" by arith
+      thus "P i" using base by simp
+    next
+      case (Suc n)
+      hence "n = nat(k - (i+1))" by arith
+      moreover
+      have ki1: "i + 1 \<le> k" using Suc.prems by arith
+      ultimately
+      have "P(i+1)" by(rule Suc.hyps)
+      from step[OF ki1 this] show ?case by simp
+    qed
+  }
+  with le show ?thesis by fast
+qed
+
+theorem int_less_induct [consumes 1,case_names base step]:
+  assumes less: "(i::int) < k" and
+        base: "P(k - 1)" and
+        step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
+  shows "P i"
+apply(rule int_le_induct[of _ "k - 1"])
+  using less apply arith
+ apply(rule base)
+apply (rule step, simp+)
+done
+
+subsection{*Intermediate value theorems*}
+
+lemma int_val_lemma:
+     "(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) -->  
+      f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))"
+apply (induct_tac "n", simp)
+apply (intro strip)
+apply (erule impE, simp)
+apply (erule_tac x = n in allE, simp)
+apply (case_tac "k = f (n+1) ")
+ apply force
+apply (erule impE)
+ apply (simp add: abs_if split add: split_if_asm)
+apply (blast intro: le_SucI)
+done
+
+lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)]
+
+lemma nat_intermed_int_val:
+     "[| \<forall>i. m \<le> i & i < n --> abs(f(i + 1::nat) - f i) \<le> 1; m < n;  
+         f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)"
+apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k 
+       in int_val_lemma)
+apply simp
+apply (erule exE)
+apply (rule_tac x = "i+m" in exI, arith)
+done
+
+
+subsection{*Products and 1, by T. M. Rasmussen*}
+
+lemma zabs_less_one_iff [simp]: "(\<bar>z\<bar> < 1) = (z = (0::int))"
+by arith
+
+lemma abs_zmult_eq_1: "(\<bar>m * n\<bar> = 1) ==> \<bar>m\<bar> = (1::int)"
+apply (cases "\<bar>n\<bar>=1") 
+apply (simp add: abs_mult) 
+apply (rule ccontr) 
+apply (auto simp add: linorder_neq_iff abs_mult) 
+apply (subgoal_tac "2 \<le> \<bar>m\<bar> & 2 \<le> \<bar>n\<bar>")
+ prefer 2 apply arith 
+apply (subgoal_tac "2*2 \<le> \<bar>m\<bar> * \<bar>n\<bar>", simp) 
+apply (rule mult_mono, auto) 
+done
+
+lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1"
+by (insert abs_zmult_eq_1 [of m n], arith)
+
+lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)"
+apply (auto dest: pos_zmult_eq_1_iff_lemma) 
+apply (simp add: mult_commute [of m]) 
+apply (frule pos_zmult_eq_1_iff_lemma, auto) 
+done
+
+lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))"
+apply (rule iffI) 
+ apply (frule pos_zmult_eq_1_iff_lemma)
+ apply (simp add: mult_commute [of m]) 
+ apply (frule pos_zmult_eq_1_iff_lemma, auto) 
+done
+
+
+subsection {* Legacy ML bindings *}
+
+ML {*
+val of_int_number_of_eq = @{thm of_int_number_of_eq};
+val nat_0 = @{thm nat_0};
+val nat_1 = @{thm nat_1};
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IntDef.thy	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,890 @@
+(*  Title:      IntDef.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1996  University of Cambridge
+
+*)
+
+header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} 
+
+theory IntDef
+imports Equiv_Relations Nat
+begin
+
+text {* the equivalence relation underlying the integers *}
+
+definition
+  intrel :: "((nat \<times> nat) \<times> (nat \<times> nat)) set"
+where
+  "intrel = {((x, y), (u, v)) | x y u v. x + v = u +y }"
+
+typedef (Integ)
+  int = "UNIV//intrel"
+  by (auto simp add: quotient_def)
+
+definition
+  int :: "nat \<Rightarrow> int"
+where
+  [code func del]: "int m = Abs_Integ (intrel `` {(m, 0)})"
+
+instance int :: zero
+  Zero_int_def: "0 \<equiv> int 0" ..
+
+instance int :: one
+  One_int_def: "1 \<equiv> int 1" ..
+
+instance int :: plus
+  add_int_def: "z + w \<equiv> Abs_Integ
+    (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u, v) \<in> Rep_Integ w.
+      intrel `` {(x + u, y + v)})" ..
+
+instance int :: minus
+  minus_int_def:
+    "- z \<equiv> Abs_Integ (\<Union>(x, y) \<in> Rep_Integ z. intrel `` {(y, x)})"
+  diff_int_def:  "z - w \<equiv> z + (-w)" ..
+
+instance int :: times
+  mult_int_def: "z * w \<equiv>  Abs_Integ
+    (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u,v ) \<in> Rep_Integ w.
+      intrel `` {(x*u + y*v, x*v + y*u)})" ..
+
+instance int :: ord
+  le_int_def:
+   "z \<le> w \<equiv> \<exists>x y u v. x+v \<le> u+y \<and> (x, y) \<in> Rep_Integ z \<and> (u, v) \<in> Rep_Integ w"
+  less_int_def: "z < w \<equiv> z \<le> w \<and> z \<noteq> w" ..
+
+lemmas [code func del] = Zero_int_def One_int_def add_int_def
+  minus_int_def mult_int_def le_int_def less_int_def
+
+
+subsection{*Construction of the Integers*}
+
+subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
+
+lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)"
+by (simp add: intrel_def)
+
+lemma equiv_intrel: "equiv UNIV intrel"
+by (simp add: intrel_def equiv_def refl_def sym_def trans_def)
+
+text{*Reduces equality of equivalence classes to the @{term intrel} relation:
+  @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
+lemmas equiv_intrel_iff [simp] = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I]
+
+text{*All equivalence classes belong to set of representatives*}
+lemma [simp]: "intrel``{(x,y)} \<in> Integ"
+by (auto simp add: Integ_def intrel_def quotient_def)
+
+text{*Reduces equality on abstractions to equality on representatives:
+  @{prop "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
+declare Abs_Integ_inject [simp]  Abs_Integ_inverse [simp]
+
+text{*Case analysis on the representation of an integer as an equivalence
+      class of pairs of naturals.*}
+lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:
+     "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
+apply (rule Abs_Integ_cases [of z]) 
+apply (auto simp add: Integ_def quotient_def) 
+done
+
+
+subsubsection{*@{term int}: Embedding the Naturals into the Integers*}
+
+lemma inj_int: "inj int"
+by (simp add: inj_on_def int_def)
+
+lemma int_int_eq [iff]: "(int m = int n) = (m = n)"
+by (fast elim!: inj_int [THEN injD])
+
+
+subsubsection{*Integer Unary Negation*}
+
+lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
+proof -
+  have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
+    by (simp add: congruent_def) 
+  thus ?thesis
+    by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel])
+qed
+
+lemma zminus_zminus: "- (- z) = (z::int)"
+  by (cases z) (simp add: minus)
+
+lemma zminus_0: "- 0 = (0::int)"
+  by (simp add: int_def Zero_int_def minus)
+
+
+subsection{*Integer Addition*}
+
+lemma add:
+     "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
+      Abs_Integ (intrel``{(x+u, y+v)})"
+proof -
+  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z) 
+        respects2 intrel"
+    by (simp add: congruent2_def)
+  thus ?thesis
+    by (simp add: add_int_def UN_UN_split_split_eq
+                  UN_equiv_class2 [OF equiv_intrel equiv_intrel])
+qed
+
+lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)"
+  by (cases z, cases w) (simp add: minus add)
+
+lemma zadd_commute: "(z::int) + w = w + z"
+  by (cases z, cases w) (simp add: add_ac add)
+
+lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
+  by (cases z1, cases z2, cases z3) (simp add: add add_assoc)
+
+(*For AC rewriting*)
+lemma zadd_left_commute: "x + (y + z) = y + ((x + z)  ::int)"
+  apply (rule mk_left_commute [of "op +"])
+  apply (rule zadd_assoc)
+  apply (rule zadd_commute)
+  done
+
+lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
+
+lemmas zmult_ac = OrderedGroup.mult_ac
+
+lemma zadd_int: "(int m) + (int n) = int (m + n)"
+  by (simp add: int_def add)
+
+lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z"
+  by (simp add: zadd_int zadd_assoc [symmetric])
+
+(*also for the instance declaration int :: comm_monoid_add*)
+lemma zadd_0: "(0::int) + z = z"
+apply (simp add: Zero_int_def int_def)
+apply (cases z, simp add: add)
+done
+
+lemma zadd_0_right: "z + (0::int) = z"
+by (rule trans [OF zadd_commute zadd_0])
+
+lemma zadd_zminus_inverse2: "(- z) + z = (0::int)"
+by (cases z, simp add: int_def Zero_int_def minus add)
+
+
+subsection{*Integer Multiplication*}
+
+text{*Congruence property for multiplication*}
+lemma mult_congruent2:
+     "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)
+      respects2 intrel"
+apply (rule equiv_intrel [THEN congruent2_commuteI])
+ apply (force simp add: mult_ac, clarify) 
+apply (simp add: congruent_def mult_ac)  
+apply (rename_tac u v w x y z)
+apply (subgoal_tac "u*y + x*y = w*y + v*y  &  u*z + x*z = w*z + v*z")
+apply (simp add: mult_ac)
+apply (simp add: add_mult_distrib [symmetric])
+done
+
+
+lemma mult:
+     "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
+      Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
+by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2
+              UN_equiv_class2 [OF equiv_intrel equiv_intrel])
+
+lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
+by (cases z, cases w, simp add: minus mult add_ac)
+
+lemma zmult_commute: "(z::int) * w = w * z"
+by (cases z, cases w, simp add: mult add_ac mult_ac)
+
+lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
+by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac)
+
+lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
+by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac)
+
+lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
+by (simp add: zmult_commute [of w] zadd_zmult_distrib)
+
+lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
+by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
+
+lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
+by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
+
+lemmas int_distrib =
+  zadd_zmult_distrib zadd_zmult_distrib2
+  zdiff_zmult_distrib zdiff_zmult_distrib2
+
+lemma int_mult: "int (m * n) = (int m) * (int n)"
+by (simp add: int_def mult)
+
+text{*Compatibility binding*}
+lemmas zmult_int = int_mult [symmetric]
+
+lemma zmult_1: "(1::int) * z = z"
+by (cases z, simp add: One_int_def int_def mult)
+
+lemma zmult_1_right: "z * (1::int) = z"
+by (rule trans [OF zmult_commute zmult_1])
+
+
+text{*The integers form a @{text comm_ring_1}*}
+instance int :: comm_ring_1
+proof
+  fix i j k :: int
+  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
+  show "i + j = j + i" by (simp add: zadd_commute)
+  show "0 + i = i" by (rule zadd_0)
+  show "- i + i = 0" by (rule zadd_zminus_inverse2)
+  show "i - j = i + (-j)" by (simp add: diff_int_def)
+  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
+  show "i * j = j * i" by (rule zmult_commute)
+  show "1 * i = i" by (rule zmult_1) 
+  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
+  show "0 \<noteq> (1::int)"
+    by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
+qed
+
+
+subsection{*The @{text "\<le>"} Ordering*}
+
+lemma le:
+  "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
+by (force simp add: le_int_def)
+
+lemma zle_refl: "w \<le> (w::int)"
+by (cases w, simp add: le)
+
+lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)"
+by (cases i, cases j, cases k, simp add: le)
+
+lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)"
+by (cases w, cases z, simp add: le)
+
+instance int :: order
+  by intro_classes
+    (assumption |
+      rule zle_refl zle_trans zle_anti_sym less_int_def [THEN meta_eq_to_obj_eq])+
+
+lemma zle_linear: "(z::int) \<le> w \<or> w \<le> z"
+by (cases z, cases w) (simp add: le linorder_linear)
+
+instance int :: linorder
+  by intro_classes (rule zle_linear)
+
+lemmas zless_linear = linorder_less_linear [where 'a = int]
+
+
+lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)"
+by (simp add: Zero_int_def)
+
+lemma zless_int [simp]: "(int m < int n) = (m<n)"
+by (simp add: le add int_def linorder_not_le [symmetric]) 
+
+lemma int_less_0_conv [simp]: "~ (int k < 0)"
+by (simp add: Zero_int_def)
+
+lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)"
+by (simp add: Zero_int_def)
+
+lemma int_0_less_1: "0 < (1::int)"
+by (simp only: Zero_int_def One_int_def One_nat_def zless_int)
+
+lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)"
+by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
+
+lemma zle_int [simp]: "(int m \<le> int n) = (m\<le>n)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma zero_zle_int [simp]: "(0 \<le> int n)"
+by (simp add: Zero_int_def)
+
+lemma int_le_0_conv [simp]: "(int n \<le> 0) = (n = 0)"
+by (simp add: Zero_int_def)
+
+lemma int_0 [simp]: "int 0 = (0::int)"
+by (simp add: Zero_int_def)
+
+lemma int_1 [simp]: "int 1 = 1"
+by (simp add: One_int_def)
+
+lemma int_Suc0_eq_1: "int (Suc 0) = 1"
+by (simp add: One_int_def One_nat_def)
+
+lemma int_Suc: "int (Suc m) = 1 + (int m)"
+by (simp add: One_int_def zadd_int)
+
+
+subsection{*Monotonicity results*}
+
+lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)"
+by (cases i, cases j, cases k, simp add: le add)
+
+lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)"
+apply (cases i, cases j, cases k)
+apply (simp add: linorder_not_le [where 'a = int, symmetric]
+                 linorder_not_le [where 'a = nat]  le add)
+done
+
+lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)"
+by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono])
+
+
+subsection{*Strict Monotonicity of Multiplication*}
+
+text{*strict, in 1st argument; proof is by induction on k>0*}
+lemma zmult_zless_mono2_lemma:
+     "i<j ==> 0<k ==> int k * i < int k * j"
+apply (induct "k", simp)
+apply (simp add: int_Suc)
+apply (case_tac "k=0")
+apply (simp_all add: zadd_zmult_distrib int_Suc0_eq_1 order_le_less)
+apply (simp add: zadd_zless_mono int_Suc0_eq_1 order_le_less)
+done
+
+lemma zero_le_imp_eq_int: "0 \<le> k ==> \<exists>n. k = int n"
+apply (cases k)
+apply (auto simp add: le add int_def Zero_int_def)
+apply (rule_tac x="x-y" in exI, simp)
+done
+
+lemma zmult_zless_mono2: "[| i<j;  (0::int) < k |] ==> k*i < k*j"
+apply (frule order_less_imp_le [THEN zero_le_imp_eq_int])
+apply (auto simp add: zmult_zless_mono2_lemma)
+done
+
+instance int :: minus
+  zabs_def: "\<bar>i\<Colon>int\<bar> \<equiv> if i < 0 then - i else i" ..
+
+instance int :: distrib_lattice
+  "inf \<equiv> min"
+  "sup \<equiv> max"
+  by intro_classes
+    (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1)
+
+text{*The integers form an ordered @{text comm_ring_1}*}
+instance int :: ordered_idom
+proof
+  fix i j k :: int
+  show "i \<le> j ==> k + i \<le> k + j" by (rule zadd_left_mono)
+  show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2)
+  show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def)
+qed
+
+lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
+apply (cases w, cases z) 
+apply (simp add: linorder_not_le [symmetric] le int_def add One_int_def)
+done
+
+subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
+
+definition
+  nat :: "int \<Rightarrow> nat"
+where
+  [code func del]: "nat z = contents (\<Union>(x, y) \<in> Rep_Integ z. {x-y})"
+
+lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
+proof -
+  have "(\<lambda>(x,y). {x-y}) respects intrel"
+    by (simp add: congruent_def) arith
+  thus ?thesis
+    by (simp add: nat_def UN_equiv_class [OF equiv_intrel])
+qed
+
+lemma nat_int [simp]: "nat(int n) = n"
+by (simp add: nat int_def) 
+
+lemma nat_zero [simp]: "nat 0 = 0"
+by (simp only: Zero_int_def nat_int)
+
+lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"
+by (cases z, simp add: nat le int_def Zero_int_def)
+
+corollary nat_0_le: "0 \<le> z ==> int (nat z) = z"
+by simp
+
+lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
+by (cases z, simp add: nat le int_def Zero_int_def)
+
+lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
+apply (cases w, cases z) 
+apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith)
+done
+
+text{*An alternative condition is @{term "0 \<le> w"} *}
+corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
+by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
+
+corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)"
+by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
+
+lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)"
+apply (cases w, cases z) 
+apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
+done
+
+lemma nonneg_eq_int: "[| 0 \<le> z;  !!m. z = int m ==> P |] ==> P"
+by (blast dest: nat_0_le sym)
+
+lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
+by (cases w, simp add: nat le int_def Zero_int_def, arith)
+
+corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
+by (simp only: eq_commute [of m] nat_eq_iff) 
+
+lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)"
+apply (cases w)
+apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
+done
+
+lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)"
+by (auto simp add: nat_eq_iff2)
+
+lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
+by (insert zless_nat_conj [of 0], auto)
+
+lemma nat_add_distrib:
+     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
+by (cases z, cases z', simp add: nat add le int_def Zero_int_def)
+
+lemma nat_diff_distrib:
+     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
+by (cases z, cases z', 
+    simp add: nat add minus diff_minus le int_def Zero_int_def)
+
+
+lemma nat_zminus_int [simp]: "nat (- (int n)) = 0"
+by (simp add: int_def minus nat Zero_int_def) 
+
+lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
+by (cases z, simp add: nat le int_def  linorder_not_le [symmetric], arith)
+
+
+subsection{*Lemmas about the Function @{term int} and Orderings*}
+
+lemma negative_zless_0: "- (int (Suc n)) < 0"
+by (simp add: order_less_le)
+
+lemma negative_zless [iff]: "- (int (Suc n)) < int m"
+by (rule negative_zless_0 [THEN order_less_le_trans], simp)
+
+lemma negative_zle_0: "- int n \<le> 0"
+by (simp add: minus_le_iff)
+
+lemma negative_zle [iff]: "- int n \<le> int m"
+by (rule order_trans [OF negative_zle_0 zero_zle_int])
+
+lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))"
+by (subst le_minus_iff, simp)
+
+lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)"
+by (simp add: int_def le minus Zero_int_def) 
+
+lemma not_int_zless_negative [simp]: "~ (int n < - int m)"
+by (simp add: linorder_not_less)
+
+lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)"
+by (force simp add: order_eq_iff [of "- int n"] int_zle_neg)
+
+lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)"
+proof (cases w, cases z, simp add: le add int_def)
+  fix a b c d
+  assume "w = Abs_Integ (intrel `` {(a,b)})" "z = Abs_Integ (intrel `` {(c,d)})"
+  show "(a+d \<le> c+b) = (\<exists>n. c+b = a+n+d)"
+  proof
+    assume "a + d \<le> c + b" 
+    thus "\<exists>n. c + b = a + n + d" 
+      by (auto intro!: exI [where x="c+b - (a+d)"])
+  next    
+    assume "\<exists>n. c + b = a + n + d" 
+    then obtain n where "c + b = a + n + d" ..
+    thus "a + d \<le> c + b" by arith
+  qed
+qed
+
+lemma abs_int_eq [simp]: "abs (int m) = int m"
+by (simp add: abs_if)
+
+text{*This version is proved for all ordered rings, not just integers!
+      It is proved here because attribute @{text arith_split} is not available
+      in theory @{text Ring_and_Field}.
+      But is it really better than just rewriting with @{text abs_if}?*}
+lemma abs_split [arith_split]:
+     "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
+by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
+
+
+subsection {* Constants @{term neg} and @{term iszero} *}
+
+definition
+  neg  :: "'a\<Colon>ordered_idom \<Rightarrow> bool"
+where
+  [code inline]: "neg Z \<longleftrightarrow> Z < 0"
+
+definition (*for simplifying equalities*)
+  iszero :: "'a\<Colon>comm_semiring_1_cancel \<Rightarrow> bool"
+where
+  "iszero z \<longleftrightarrow> z = 0"
+
+lemma not_neg_int [simp]: "~ neg(int n)"
+by (simp add: neg_def)
+
+lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))"
+by (simp add: neg_def neg_less_0_iff_less)
+
+lemmas neg_eq_less_0 = neg_def
+
+lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
+by (simp add: neg_def linorder_not_less)
+
+
+subsection{*To simplify inequalities when Numeral1 can get simplified to 1*}
+
+lemma not_neg_0: "~ neg 0"
+by (simp add: One_int_def neg_def)
+
+lemma not_neg_1: "~ neg 1"
+by (simp add: neg_def linorder_not_less zero_le_one)
+
+lemma iszero_0: "iszero 0"
+by (simp add: iszero_def)
+
+lemma not_iszero_1: "~ iszero 1"
+by (simp add: iszero_def eq_commute)
+
+lemma neg_nat: "neg z ==> nat z = 0"
+by (simp add: neg_def order_less_imp_le) 
+
+lemma not_neg_nat: "~ neg z ==> int (nat z) = z"
+by (simp add: linorder_not_less neg_def)
+
+
+subsection{*The Set of Natural Numbers*}
+
+constdefs
+  Nats  :: "'a::semiring_1_cancel set"
+  "Nats == range of_nat"
+
+notation (xsymbols)
+  Nats  ("\<nat>")
+
+lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats"
+by (simp add: Nats_def)
+
+lemma Nats_0 [simp]: "0 \<in> Nats"
+apply (simp add: Nats_def)
+apply (rule range_eqI)
+apply (rule of_nat_0 [symmetric])
+done
+
+lemma Nats_1 [simp]: "1 \<in> Nats"
+apply (simp add: Nats_def)
+apply (rule range_eqI)
+apply (rule of_nat_1 [symmetric])
+done
+
+lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats"
+apply (auto simp add: Nats_def)
+apply (rule range_eqI)
+apply (rule of_nat_add [symmetric])
+done
+
+lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats"
+apply (auto simp add: Nats_def)
+apply (rule range_eqI)
+apply (rule of_nat_mult [symmetric])
+done
+
+text{*Agreement with the specific embedding for the integers*}
+lemma int_eq_of_nat: "int = (of_nat :: nat => int)"
+proof
+  fix n
+  show "int n = of_nat n"  by (induct n, simp_all add: int_Suc add_ac)
+qed
+
+lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)"
+proof
+  fix n
+  show "of_nat n = id n"  by (induct n, simp_all)
+qed
+
+
+subsection{*Embedding of the Integers into any @{text ring_1}:
+@{term of_int}*}
+
+constdefs
+   of_int :: "int => 'a::ring_1"
+   "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
+
+
+lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
+proof -
+  have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
+    by (simp add: congruent_def compare_rls of_nat_add [symmetric]
+            del: of_nat_add) 
+  thus ?thesis
+    by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
+qed
+
+lemma of_int_0 [simp]: "of_int 0 = 0"
+by (simp add: of_int Zero_int_def int_def)
+
+lemma of_int_1 [simp]: "of_int 1 = 1"
+by (simp add: of_int One_int_def int_def)
+
+lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z"
+by (cases w, cases z, simp add: compare_rls of_int add)
+
+lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
+by (cases z, simp add: compare_rls of_int minus)
+
+lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z"
+by (simp add: diff_minus)
+
+lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
+apply (cases w, cases z)
+apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
+                 mult add_ac)
+done
+
+lemma of_int_le_iff [simp]:
+     "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
+apply (cases w)
+apply (cases z)
+apply (simp add: compare_rls of_int le diff_int_def add minus
+                 of_nat_add [symmetric]   del: of_nat_add)
+done
+
+text{*Special cases where either operand is zero*}
+lemmas of_int_0_le_iff [simp] = of_int_le_iff [of 0, simplified]
+lemmas of_int_le_0_iff [simp] = of_int_le_iff [of _ 0, simplified]
+
+
+lemma of_int_less_iff [simp]:
+     "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
+by (simp add: linorder_not_le [symmetric])
+
+text{*Special cases where either operand is zero*}
+lemmas of_int_0_less_iff [simp] = of_int_less_iff [of 0, simplified]
+lemmas of_int_less_0_iff [simp] = of_int_less_iff [of _ 0, simplified]
+
+text{*Class for unital rings with characteristic zero.
+ Includes non-ordered rings like the complex numbers.*}
+axclass ring_char_0 < ring_1
+  of_int_inject: "of_int w = of_int z ==> w = z"
+
+lemma of_int_eq_iff [simp]:
+     "(of_int w = (of_int z::'a::ring_char_0)) = (w = z)"
+by (safe elim!: of_int_inject)
+
+text{*Every @{text ordered_idom} has characteristic zero.*}
+instance ordered_idom < ring_char_0
+by intro_classes (simp add: order_eq_iff)
+
+text{*Special cases where either operand is zero*}
+lemmas of_int_0_eq_iff [simp] = of_int_eq_iff [of 0, simplified]
+lemmas of_int_eq_0_iff [simp] = of_int_eq_iff [of _ 0, simplified]
+
+lemma of_int_eq_id [simp]: "of_int = (id :: int => int)"
+proof
+  fix z
+  show "of_int z = id z"  
+    by (cases z)
+      (simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus)
+qed
+
+
+subsection{*The Set of Integers*}
+
+constdefs
+  Ints  :: "'a::ring_1 set"
+  "Ints == range of_int"
+
+notation (xsymbols)
+  Ints  ("\<int>")
+
+lemma Ints_0 [simp]: "0 \<in> Ints"
+apply (simp add: Ints_def)
+apply (rule range_eqI)
+apply (rule of_int_0 [symmetric])
+done
+
+lemma Ints_1 [simp]: "1 \<in> Ints"
+apply (simp add: Ints_def)
+apply (rule range_eqI)
+apply (rule of_int_1 [symmetric])
+done
+
+lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints"
+apply (auto simp add: Ints_def)
+apply (rule range_eqI)
+apply (rule of_int_add [symmetric])
+done
+
+lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints"
+apply (auto simp add: Ints_def)
+apply (rule range_eqI)
+apply (rule of_int_minus [symmetric])
+done
+
+lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints"
+apply (auto simp add: Ints_def)
+apply (rule range_eqI)
+apply (rule of_int_diff [symmetric])
+done
+
+lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints"
+apply (auto simp add: Ints_def)
+apply (rule range_eqI)
+apply (rule of_int_mult [symmetric])
+done
+
+text{*Collapse nested embeddings*}
+lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
+by (induct n, auto)
+
+lemma of_int_int_eq [simp]: "of_int (int n) = of_nat n"
+by (simp add: int_eq_of_nat)
+
+lemma Ints_cases [cases set: Ints]:
+  assumes "q \<in> \<int>"
+  obtains (of_int) z where "q = of_int z"
+  unfolding Ints_def
+proof -
+  from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def .
+  then obtain z where "q = of_int z" ..
+  then show thesis ..
+qed
+
+lemma Ints_induct [case_names of_int, induct set: Ints]:
+  "q \<in> \<int> ==> (!!z. P (of_int z)) ==> P q"
+  by (rule Ints_cases) auto
+
+
+(* int (Suc n) = 1 + int n *)
+
+
+
+subsection{*More Properties of @{term setsum} and  @{term setprod}*}
+
+text{*By Jeremy Avigad*}
+
+
+lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
+  apply (cases "finite A")
+  apply (erule finite_induct, auto)
+  done
+
+lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))"
+  apply (cases "finite A")
+  apply (erule finite_induct, auto)
+  done
+
+lemma int_setsum: "int (setsum f A) = (\<Sum>x\<in>A. int(f x))"
+  by (simp add: int_eq_of_nat of_nat_setsum)
+
+lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))"
+  apply (cases "finite A")
+  apply (erule finite_induct, auto)
+  done
+
+lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
+  apply (cases "finite A")
+  apply (erule finite_induct, auto)
+  done
+
+lemma int_setprod: "int (setprod f A) = (\<Prod>x\<in>A. int(f x))"
+  by (simp add: int_eq_of_nat of_nat_setprod)
+
+lemma setprod_nonzero_nat:
+    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
+  by (rule setprod_nonzero, auto)
+
+lemma setprod_zero_eq_nat:
+    "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
+  by (rule setprod_zero_eq, auto)
+
+lemma setprod_nonzero_int:
+    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
+  by (rule setprod_nonzero, auto)
+
+lemma setprod_zero_eq_int:
+    "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
+  by (rule setprod_zero_eq, auto)
+
+
+subsection {* Further properties *}
+
+text{*Now we replace the case analysis rule by a more conventional one:
+whether an integer is negative or not.*}
+
+lemma zless_iff_Suc_zadd:
+    "(w < z) = (\<exists>n. z = w + int(Suc n))"
+apply (cases z, cases w)
+apply (auto simp add: le add int_def linorder_not_le [symmetric]) 
+apply (rename_tac a b c d) 
+apply (rule_tac x="a+d - Suc(c+b)" in exI) 
+apply arith
+done
+
+lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
+apply (cases x)
+apply (auto simp add: le minus Zero_int_def int_def order_less_le) 
+apply (rule_tac x="y - Suc x" in exI, arith)
+done
+
+theorem int_cases [cases type: int, case_names nonneg neg]:
+     "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
+apply (cases "z < 0", blast dest!: negD)
+apply (simp add: linorder_not_less)
+apply (blast dest: nat_0_le [THEN sym])
+done
+
+theorem int_induct [induct type: int, case_names nonneg neg]:
+     "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
+  by (cases z) auto
+
+text{*Contributed by Brian Huffman*}
+theorem int_diff_cases [case_names diff]:
+assumes prem: "!!m n. z = int m - int n ==> P" shows "P"
+ apply (rule_tac z=z in int_cases)
+  apply (rule_tac m=n and n=0 in prem, simp)
+ apply (rule_tac m=0 and n="Suc n" in prem, simp)
+done
+
+lemma of_nat_nat: "0 \<le> z ==> of_nat (nat z) = of_int z"
+apply (cases z)
+apply (simp_all add: not_zle_0_negative del: int_Suc)
+done
+
+lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq]
+
+lemmas [simp] = int_Suc
+
+
+subsection {* Legacy ML bindings *}
+
+ML {*
+val of_nat_0 = @{thm of_nat_0};
+val of_nat_1 = @{thm of_nat_1};
+val of_nat_Suc = @{thm of_nat_Suc};
+val of_nat_add = @{thm of_nat_add};
+val of_nat_mult = @{thm of_nat_mult};
+val of_int_0 = @{thm of_int_0};
+val of_int_1 = @{thm of_int_1};
+val of_int_add = @{thm of_int_add};
+val of_int_mult = @{thm of_int_mult};
+val int_eq_of_nat = @{thm int_eq_of_nat};
+val zle_int = @{thm zle_int};
+val int_int_eq = @{thm int_int_eq};
+val diff_int_def = @{thm diff_int_def};
+val zadd_ac = @{thms zadd_ac};
+val zless_int = @{thm zless_int};
+val zadd_int = @{thm zadd_int};
+val zmult_int = @{thm zmult_int};
+val nat_0_le = @{thm nat_0_le};
+val int_0 = @{thm int_0};
+val int_1 = @{thm int_1};
+val abs_split = @{thm abs_split};
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/IntDiv.thy	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,1406 @@
+(*  Title:      HOL/IntDiv.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+
+*)
+
+header{*The Division Operators div and mod; the Divides Relation dvd*}
+
+theory IntDiv
+imports IntArith Divides FunDef
+begin
+
+declare zless_nat_conj [simp]
+
+constdefs
+  quorem :: "(int*int) * (int*int) => bool"
+    --{*definition of quotient and remainder*}
+    [code func]: "quorem == %((a,b), (q,r)).
+                      a = b*q + r &
+                      (if 0 < b then 0\<le>r & r<b else b<r & r \<le> 0)"
+
+  adjust :: "[int, int*int] => int*int"
+    --{*for the division algorithm*}
+    [code func]: "adjust b == %(q,r). if 0 \<le> r-b then (2*q + 1, r-b)
+                         else (2*q, r)"
+
+text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
+function
+  posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
+where
+  "posDivAlg a b =
+     (if (a<b | b\<le>0) then (0,a)
+        else adjust b (posDivAlg a (2*b)))"
+by auto
+termination by (relation "measure (%(a,b). nat(a - b + 1))") auto
+
+text{*algorithm for the case @{text "a<0, b>0"}*}
+function
+  negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
+where
+  "negDivAlg a b  =
+     (if (0\<le>a+b | b\<le>0) then (-1,a+b)
+      else adjust b (negDivAlg a (2*b)))"
+by auto
+termination by (relation "measure (%(a,b). nat(- a - b))") auto
+
+text{*algorithm for the general case @{term "b\<noteq>0"}*}
+constdefs
+  negateSnd :: "int*int => int*int"
+    [code func]: "negateSnd == %(q,r). (q,-r)"
+
+definition
+  divAlg :: "int \<times> int \<Rightarrow> int \<times> int"
+    --{*The full division algorithm considers all possible signs for a, b
+       including the special case @{text "a=0, b<0"} because 
+       @{term negDivAlg} requires @{term "a<0"}.*}
+where
+  "divAlg = (\<lambda>(a, b). (if 0\<le>a then
+                  if 0\<le>b then posDivAlg a b
+                  else if a=0 then (0, 0)
+                       else negateSnd (negDivAlg (-a) (-b))
+               else 
+                  if 0<b then negDivAlg a b
+                  else negateSnd (posDivAlg (-a) (-b))))"
+
+instance int :: Divides.div
+  div_def: "a div b == fst (divAlg (a, b))"
+  mod_def: "a mod b == snd (divAlg (a, b))" ..
+
+lemma divAlg_mod_div:
+  "divAlg (p, q) = (p div q, p mod q)"
+  by (auto simp add: div_def mod_def)
+
+text{*
+Here is the division algorithm in ML:
+
+\begin{verbatim}
+    fun posDivAlg (a,b) =
+      if a<b then (0,a)
+      else let val (q,r) = posDivAlg(a, 2*b)
+	       in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
+	   end
+
+    fun negDivAlg (a,b) =
+      if 0\<le>a+b then (~1,a+b)
+      else let val (q,r) = negDivAlg(a, 2*b)
+	       in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
+	   end;
+
+    fun negateSnd (q,r:int) = (q,~r);
+
+    fun divAlg (a,b) = if 0\<le>a then 
+			  if b>0 then posDivAlg (a,b) 
+			   else if a=0 then (0,0)
+				else negateSnd (negDivAlg (~a,~b))
+		       else 
+			  if 0<b then negDivAlg (a,b)
+			  else        negateSnd (posDivAlg (~a,~b));
+\end{verbatim}
+*}
+
+
+
+subsection{*Uniqueness and Monotonicity of Quotients and Remainders*}
+
+lemma unique_quotient_lemma:
+     "[| b*q' + r'  \<le> b*q + r;  0 \<le> r';  r' < b;  r < b |]  
+      ==> q' \<le> (q::int)"
+apply (subgoal_tac "r' + b * (q'-q) \<le> r")
+ prefer 2 apply (simp add: right_diff_distrib)
+apply (subgoal_tac "0 < b * (1 + q - q') ")
+apply (erule_tac [2] order_le_less_trans)
+ prefer 2 apply (simp add: right_diff_distrib right_distrib)
+apply (subgoal_tac "b * q' < b * (1 + q) ")
+ prefer 2 apply (simp add: right_diff_distrib right_distrib)
+apply (simp add: mult_less_cancel_left)
+done
+
+lemma unique_quotient_lemma_neg:
+     "[| b*q' + r' \<le> b*q + r;  r \<le> 0;  b < r;  b < r' |]  
+      ==> q \<le> (q'::int)"
+by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, 
+    auto)
+
+lemma unique_quotient:
+     "[| quorem ((a,b), (q,r));  quorem ((a,b), (q',r'));  b \<noteq> 0 |]  
+      ==> q = q'"
+apply (simp add: quorem_def linorder_neq_iff split: split_if_asm)
+apply (blast intro: order_antisym
+             dest: order_eq_refl [THEN unique_quotient_lemma] 
+             order_eq_refl [THEN unique_quotient_lemma_neg] sym)+
+done
+
+
+lemma unique_remainder:
+     "[| quorem ((a,b), (q,r));  quorem ((a,b), (q',r'));  b \<noteq> 0 |]  
+      ==> r = r'"
+apply (subgoal_tac "q = q'")
+ apply (simp add: quorem_def)
+apply (blast intro: unique_quotient)
+done
+
+
+subsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*}
+
+text{*And positive divisors*}
+
+lemma adjust_eq [simp]:
+     "adjust b (q,r) = 
+      (let diff = r-b in  
+	if 0 \<le> diff then (2*q + 1, diff)   
+                     else (2*q, r))"
+by (simp add: Let_def adjust_def)
+
+declare posDivAlg.simps [simp del]
+
+text{*use with a simproc to avoid repeatedly proving the premise*}
+lemma posDivAlg_eqn:
+     "0 < b ==>  
+      posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))"
+by (rule posDivAlg.simps [THEN trans], simp)
+
+text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
+theorem posDivAlg_correct:
+  assumes "0 \<le> a" and "0 < b"
+  shows "quorem ((a, b), posDivAlg a b)"
+using prems apply (induct a b rule: posDivAlg.induct)
+apply auto
+apply (simp add: quorem_def)
+apply (subst posDivAlg_eqn, simp add: right_distrib)
+apply (case_tac "a < b")
+apply simp_all
+apply (erule splitE)
+apply (auto simp add: right_distrib Let_def)
+done
+
+
+subsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*}
+
+text{*And positive divisors*}
+
+declare negDivAlg.simps [simp del]
+
+text{*use with a simproc to avoid repeatedly proving the premise*}
+lemma negDivAlg_eqn:
+     "0 < b ==>  
+      negDivAlg a b =       
+       (if 0\<le>a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))"
+by (rule negDivAlg.simps [THEN trans], simp)
+
+(*Correctness of negDivAlg: it computes quotients correctly
+  It doesn't work if a=0 because the 0/b equals 0, not -1*)
+lemma negDivAlg_correct:
+  assumes "a < 0" and "b > 0"
+  shows "quorem ((a, b), negDivAlg a b)"
+using prems apply (induct a b rule: negDivAlg.induct)
+apply (auto simp add: linorder_not_le)
+apply (simp add: quorem_def)
+apply (subst negDivAlg_eqn, assumption)
+apply (case_tac "a + b < (0\<Colon>int)")
+apply simp_all
+apply (erule splitE)
+apply (auto simp add: right_distrib Let_def)
+done
+
+
+subsection{*Existence Shown by Proving the Division Algorithm to be Correct*}
+
+(*the case a=0*)
+lemma quorem_0: "b \<noteq> 0 ==> quorem ((0,b), (0,0))"
+by (auto simp add: quorem_def linorder_neq_iff)
+
+lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)"
+by (subst posDivAlg.simps, auto)
+
+lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)"
+by (subst negDivAlg.simps, auto)
+
+lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)"
+by (simp add: negateSnd_def)
+
+lemma quorem_neg: "quorem ((-a,-b), qr) ==> quorem ((a,b), negateSnd qr)"
+by (auto simp add: split_ifs quorem_def)
+
+lemma divAlg_correct: "b \<noteq> 0 ==> quorem ((a,b), divAlg (a, b))"
+by (force simp add: linorder_neq_iff quorem_0 divAlg_def quorem_neg
+                    posDivAlg_correct negDivAlg_correct)
+
+text{*Arbitrary definitions for division by zero.  Useful to simplify 
+    certain equations.*}
+
+lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a"
+by (simp add: div_def mod_def divAlg_def posDivAlg.simps)  
+
+
+text{*Basic laws about division and remainder*}
+
+lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)"
+apply (case_tac "b = 0", simp)
+apply (cut_tac a = a and b = b in divAlg_correct)
+apply (auto simp add: quorem_def div_def mod_def)
+done
+
+lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k"
+by(simp add: zmod_zdiv_equality[symmetric])
+
+lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k"
+by(simp add: mult_commute zmod_zdiv_equality[symmetric])
+
+text {* Tool setup *}
+
+ML_setup {*
+local 
+
+structure CancelDivMod = CancelDivModFun(
+struct
+  val div_name = @{const_name Divides.div};
+  val mod_name = @{const_name Divides.mod};
+  val mk_binop = HOLogic.mk_binop;
+  val mk_sum = Int_Numeral_Simprocs.mk_sum HOLogic.intT;
+  val dest_sum = Int_Numeral_Simprocs.dest_sum;
+  val div_mod_eqs =
+    map mk_meta_eq [@{thm zdiv_zmod_equality},
+      @{thm zdiv_zmod_equality2}];
+  val trans = trans;
+  val prove_eq_sums =
+    let
+      val simps = diff_int_def :: Int_Numeral_Simprocs.add_0s @ zadd_ac
+    in NatArithUtils.prove_conv all_tac (NatArithUtils.simp_all_tac simps) end;
+end)
+
+in
+
+val cancel_zdiv_zmod_proc = NatArithUtils.prep_simproc
+  ("cancel_zdiv_zmod", ["(m::int) + n"], K CancelDivMod.proc)
+
+end;
+
+Addsimprocs [cancel_zdiv_zmod_proc]
+*}
+
+lemma pos_mod_conj : "(0::int) < b ==> 0 \<le> a mod b & a mod b < b"
+apply (cut_tac a = a and b = b in divAlg_correct)
+apply (auto simp add: quorem_def mod_def)
+done
+
+lemmas pos_mod_sign  [simp] = pos_mod_conj [THEN conjunct1, standard]
+   and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard]
+
+lemma neg_mod_conj : "b < (0::int) ==> a mod b \<le> 0 & b < a mod b"
+apply (cut_tac a = a and b = b in divAlg_correct)
+apply (auto simp add: quorem_def div_def mod_def)
+done
+
+lemmas neg_mod_sign  [simp] = neg_mod_conj [THEN conjunct1, standard]
+   and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard]
+
+
+
+subsection{*General Properties of div and mod*}
+
+lemma quorem_div_mod: "b \<noteq> 0 ==> quorem ((a, b), (a div b, a mod b))"
+apply (cut_tac a = a and b = b in zmod_zdiv_equality)
+apply (force simp add: quorem_def linorder_neq_iff)
+done
+
+lemma quorem_div: "[| quorem((a,b),(q,r));  b \<noteq> 0 |] ==> a div b = q"
+by (simp add: quorem_div_mod [THEN unique_quotient])
+
+lemma quorem_mod: "[| quorem((a,b),(q,r));  b \<noteq> 0 |] ==> a mod b = r"
+by (simp add: quorem_div_mod [THEN unique_remainder])
+
+lemma div_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a div b = 0"
+apply (rule quorem_div)
+apply (auto simp add: quorem_def)
+done
+
+lemma div_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a div b = 0"
+apply (rule quorem_div)
+apply (auto simp add: quorem_def)
+done
+
+lemma div_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a div b = -1"
+apply (rule quorem_div)
+apply (auto simp add: quorem_def)
+done
+
+(*There is no div_neg_pos_trivial because  0 div b = 0 would supersede it*)
+
+lemma mod_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a mod b = a"
+apply (rule_tac q = 0 in quorem_mod)
+apply (auto simp add: quorem_def)
+done
+
+lemma mod_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a mod b = a"
+apply (rule_tac q = 0 in quorem_mod)
+apply (auto simp add: quorem_def)
+done
+
+lemma mod_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a mod b = a+b"
+apply (rule_tac q = "-1" in quorem_mod)
+apply (auto simp add: quorem_def)
+done
+
+text{*There is no @{text mod_neg_pos_trivial}.*}
+
+
+(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*)
+lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)"
+apply (case_tac "b = 0", simp)
+apply (simp add: quorem_div_mod [THEN quorem_neg, simplified, 
+                                 THEN quorem_div, THEN sym])
+
+done
+
+(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*)
+lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))"
+apply (case_tac "b = 0", simp)
+apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod],
+       auto)
+done
+
+
+subsection{*Laws for div and mod with Unary Minus*}
+
+lemma zminus1_lemma:
+     "quorem((a,b),(q,r))  
+      ==> quorem ((-a,b), (if r=0 then -q else -q - 1),  
+                          (if r=0 then 0 else b-r))"
+by (force simp add: split_ifs quorem_def linorder_neq_iff right_diff_distrib)
+
+
+lemma zdiv_zminus1_eq_if:
+     "b \<noteq> (0::int)  
+      ==> (-a) div b =  
+          (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
+by (blast intro: quorem_div_mod [THEN zminus1_lemma, THEN quorem_div])
+
+lemma zmod_zminus1_eq_if:
+     "(-a::int) mod b = (if a mod b = 0 then 0 else  b - (a mod b))"
+apply (case_tac "b = 0", simp)
+apply (blast intro: quorem_div_mod [THEN zminus1_lemma, THEN quorem_mod])
+done
+
+lemma zdiv_zminus2: "a div (-b) = (-a::int) div b"
+by (cut_tac a = "-a" in zdiv_zminus_zminus, auto)
+
+lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)"
+by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto)
+
+lemma zdiv_zminus2_eq_if:
+     "b \<noteq> (0::int)  
+      ==> a div (-b) =  
+          (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
+by (simp add: zdiv_zminus1_eq_if zdiv_zminus2)
+
+lemma zmod_zminus2_eq_if:
+     "a mod (-b::int) = (if a mod b = 0 then 0 else  (a mod b) - b)"
+by (simp add: zmod_zminus1_eq_if zmod_zminus2)
+
+
+subsection{*Division of a Number by Itself*}
+
+lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \<le> q"
+apply (subgoal_tac "0 < a*q")
+ apply (simp add: zero_less_mult_iff, arith)
+done
+
+lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \<le> r |] ==> q \<le> 1"
+apply (subgoal_tac "0 \<le> a* (1-q) ")
+ apply (simp add: zero_le_mult_iff)
+apply (simp add: right_diff_distrib)
+done
+
+lemma self_quotient: "[| quorem((a,a),(q,r));  a \<noteq> (0::int) |] ==> q = 1"
+apply (simp add: split_ifs quorem_def linorder_neq_iff)
+apply (rule order_antisym, safe, simp_all)
+apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1)
+apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2)
+apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+
+done
+
+lemma self_remainder: "[| quorem((a,a),(q,r));  a \<noteq> (0::int) |] ==> r = 0"
+apply (frule self_quotient, assumption)
+apply (simp add: quorem_def)
+done
+
+lemma zdiv_self [simp]: "a \<noteq> 0 ==> a div a = (1::int)"
+by (simp add: quorem_div_mod [THEN self_quotient])
+
+(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *)
+lemma zmod_self [simp]: "a mod a = (0::int)"
+apply (case_tac "a = 0", simp)
+apply (simp add: quorem_div_mod [THEN self_remainder])
+done
+
+
+subsection{*Computation of Division and Remainder*}
+
+lemma zdiv_zero [simp]: "(0::int) div b = 0"
+by (simp add: div_def divAlg_def)
+
+lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1"
+by (simp add: div_def divAlg_def)
+
+lemma zmod_zero [simp]: "(0::int) mod b = 0"
+by (simp add: mod_def divAlg_def)
+
+lemma zdiv_minus1: "(0::int) < b ==> -1 div b = -1"
+by (simp add: div_def divAlg_def)
+
+lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1"
+by (simp add: mod_def divAlg_def)
+
+text{*a positive, b positive *}
+
+lemma div_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a div b = fst (posDivAlg a b)"
+by (simp add: div_def divAlg_def)
+
+lemma mod_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a mod b = snd (posDivAlg a b)"
+by (simp add: mod_def divAlg_def)
+
+text{*a negative, b positive *}
+
+lemma div_neg_pos: "[| a < 0;  0 < b |] ==> a div b = fst (negDivAlg a b)"
+by (simp add: div_def divAlg_def)
+
+lemma mod_neg_pos: "[| a < 0;  0 < b |] ==> a mod b = snd (negDivAlg a b)"
+by (simp add: mod_def divAlg_def)
+
+text{*a positive, b negative *}
+
+lemma div_pos_neg:
+     "[| 0 < a;  b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))"
+by (simp add: div_def divAlg_def)
+
+lemma mod_pos_neg:
+     "[| 0 < a;  b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))"
+by (simp add: mod_def divAlg_def)
+
+text{*a negative, b negative *}
+
+lemma div_neg_neg:
+     "[| a < 0;  b \<le> 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))"
+by (simp add: div_def divAlg_def)
+
+lemma mod_neg_neg:
+     "[| a < 0;  b \<le> 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))"
+by (simp add: mod_def divAlg_def)
+
+text {*Simplify expresions in which div and mod combine numerical constants*}
+
+lemmas div_pos_pos_number_of [simp] =
+    div_pos_pos [of "number_of v" "number_of w", standard]
+
+lemmas div_neg_pos_number_of [simp] =
+    div_neg_pos [of "number_of v" "number_of w", standard]
+
+lemmas div_pos_neg_number_of [simp] =
+    div_pos_neg [of "number_of v" "number_of w", standard]
+
+lemmas div_neg_neg_number_of [simp] =
+    div_neg_neg [of "number_of v" "number_of w", standard]
+
+
+lemmas mod_pos_pos_number_of [simp] =
+    mod_pos_pos [of "number_of v" "number_of w", standard]
+
+lemmas mod_neg_pos_number_of [simp] =
+    mod_neg_pos [of "number_of v" "number_of w", standard]
+
+lemmas mod_pos_neg_number_of [simp] =
+    mod_pos_neg [of "number_of v" "number_of w", standard]
+
+lemmas mod_neg_neg_number_of [simp] =
+    mod_neg_neg [of "number_of v" "number_of w", standard]
+
+
+lemmas posDivAlg_eqn_number_of [simp] =
+    posDivAlg_eqn [of "number_of v" "number_of w", standard]
+
+lemmas negDivAlg_eqn_number_of [simp] =
+    negDivAlg_eqn [of "number_of v" "number_of w", standard]
+
+
+text{*Special-case simplification *}
+
+lemma zmod_1 [simp]: "a mod (1::int) = 0"
+apply (cut_tac a = a and b = 1 in pos_mod_sign)
+apply (cut_tac [2] a = a and b = 1 in pos_mod_bound)
+apply (auto simp del:pos_mod_bound pos_mod_sign)
+done
+
+lemma zdiv_1 [simp]: "a div (1::int) = a"
+by (cut_tac a = a and b = 1 in zmod_zdiv_equality, auto)
+
+lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0"
+apply (cut_tac a = a and b = "-1" in neg_mod_sign)
+apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound)
+apply (auto simp del: neg_mod_sign neg_mod_bound)
+done
+
+lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a"
+by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto)
+
+(** The last remaining special cases for constant arithmetic:
+    1 div z and 1 mod z **)
+
+lemmas div_pos_pos_1_number_of [simp] =
+    div_pos_pos [OF int_0_less_1, of "number_of w", standard]
+
+lemmas div_pos_neg_1_number_of [simp] =
+    div_pos_neg [OF int_0_less_1, of "number_of w", standard]
+
+lemmas mod_pos_pos_1_number_of [simp] =
+    mod_pos_pos [OF int_0_less_1, of "number_of w", standard]
+
+lemmas mod_pos_neg_1_number_of [simp] =
+    mod_pos_neg [OF int_0_less_1, of "number_of w", standard]
+
+
+lemmas posDivAlg_eqn_1_number_of [simp] =
+    posDivAlg_eqn [of concl: 1 "number_of w", standard]
+
+lemmas negDivAlg_eqn_1_number_of [simp] =
+    negDivAlg_eqn [of concl: 1 "number_of w", standard]
+
+
+
+subsection{*Monotonicity in the First Argument (Dividend)*}
+
+lemma zdiv_mono1: "[| a \<le> a';  0 < (b::int) |] ==> a div b \<le> a' div b"
+apply (cut_tac a = a and b = b in zmod_zdiv_equality)
+apply (cut_tac a = a' and b = b in zmod_zdiv_equality)
+apply (rule unique_quotient_lemma)
+apply (erule subst)
+apply (erule subst, simp_all)
+done
+
+lemma zdiv_mono1_neg: "[| a \<le> a';  (b::int) < 0 |] ==> a' div b \<le> a div b"
+apply (cut_tac a = a and b = b in zmod_zdiv_equality)
+apply (cut_tac a = a' and b = b in zmod_zdiv_equality)
+apply (rule unique_quotient_lemma_neg)
+apply (erule subst)
+apply (erule subst, simp_all)
+done
+
+
+subsection{*Monotonicity in the Second Argument (Divisor)*}
+
+lemma q_pos_lemma:
+     "[| 0 \<le> b'*q' + r'; r' < b';  0 < b' |] ==> 0 \<le> (q'::int)"
+apply (subgoal_tac "0 < b'* (q' + 1) ")
+ apply (simp add: zero_less_mult_iff)
+apply (simp add: right_distrib)
+done
+
+lemma zdiv_mono2_lemma:
+     "[| b*q + r = b'*q' + r';  0 \<le> b'*q' + r';   
+         r' < b';  0 \<le> r;  0 < b';  b' \<le> b |]   
+      ==> q \<le> (q'::int)"
+apply (frule q_pos_lemma, assumption+) 
+apply (subgoal_tac "b*q < b* (q' + 1) ")
+ apply (simp add: mult_less_cancel_left)
+apply (subgoal_tac "b*q = r' - r + b'*q'")
+ prefer 2 apply simp
+apply (simp (no_asm_simp) add: right_distrib)
+apply (subst add_commute, rule zadd_zless_mono, arith)
+apply (rule mult_right_mono, auto)
+done
+
+lemma zdiv_mono2:
+     "[| (0::int) \<le> a;  0 < b';  b' \<le> b |] ==> a div b \<le> a div b'"
+apply (subgoal_tac "b \<noteq> 0")
+ prefer 2 apply arith
+apply (cut_tac a = a and b = b in zmod_zdiv_equality)
+apply (cut_tac a = a and b = b' in zmod_zdiv_equality)
+apply (rule zdiv_mono2_lemma)
+apply (erule subst)
+apply (erule subst, simp_all)
+done
+
+lemma q_neg_lemma:
+     "[| b'*q' + r' < 0;  0 \<le> r';  0 < b' |] ==> q' \<le> (0::int)"
+apply (subgoal_tac "b'*q' < 0")
+ apply (simp add: mult_less_0_iff, arith)
+done
+
+lemma zdiv_mono2_neg_lemma:
+     "[| b*q + r = b'*q' + r';  b'*q' + r' < 0;   
+         r < b;  0 \<le> r';  0 < b';  b' \<le> b |]   
+      ==> q' \<le> (q::int)"
+apply (frule q_neg_lemma, assumption+) 
+apply (subgoal_tac "b*q' < b* (q + 1) ")
+ apply (simp add: mult_less_cancel_left)
+apply (simp add: right_distrib)
+apply (subgoal_tac "b*q' \<le> b'*q'")
+ prefer 2 apply (simp add: mult_right_mono_neg, arith)
+done
+
+lemma zdiv_mono2_neg:
+     "[| a < (0::int);  0 < b';  b' \<le> b |] ==> a div b' \<le> a div b"
+apply (cut_tac a = a and b = b in zmod_zdiv_equality)
+apply (cut_tac a = a and b = b' in zmod_zdiv_equality)
+apply (rule zdiv_mono2_neg_lemma)
+apply (erule subst)
+apply (erule subst, simp_all)
+done
+
+subsection{*More Algebraic Laws for div and mod*}
+
+text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
+
+lemma zmult1_lemma:
+     "[| quorem((b,c),(q,r));  c \<noteq> 0 |]  
+      ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))"
+by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib)
+
+lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"
+apply (case_tac "c = 0", simp)
+apply (blast intro: quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
+done
+
+lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)"
+apply (case_tac "c = 0", simp)
+apply (blast intro: quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
+done
+
+lemma zmod_zmult1_eq': "(a*b) mod (c::int) = ((a mod c) * b) mod c"
+apply (rule trans)
+apply (rule_tac s = "b*a mod c" in trans)
+apply (rule_tac [2] zmod_zmult1_eq)
+apply (simp_all add: mult_commute)
+done
+
+lemma zmod_zmult_distrib: "(a*b) mod (c::int) = ((a mod c) * (b mod c)) mod c"
+apply (rule zmod_zmult1_eq' [THEN trans])
+apply (rule zmod_zmult1_eq)
+done
+
+lemma zdiv_zmult_self1 [simp]: "b \<noteq> (0::int) ==> (a*b) div b = a"
+by (simp add: zdiv_zmult1_eq)
+
+lemma zdiv_zmult_self2 [simp]: "b \<noteq> (0::int) ==> (b*a) div b = a"
+by (subst mult_commute, erule zdiv_zmult_self1)
+
+lemma zmod_zmult_self1 [simp]: "(a*b) mod b = (0::int)"
+by (simp add: zmod_zmult1_eq)
+
+lemma zmod_zmult_self2 [simp]: "(b*a) mod b = (0::int)"
+by (simp add: mult_commute zmod_zmult1_eq)
+
+lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
+proof
+  assume "m mod d = 0"
+  with zmod_zdiv_equality[of m d] show "EX q::int. m = d*q" by auto
+next
+  assume "EX q::int. m = d*q"
+  thus "m mod d = 0" by auto
+qed
+
+lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1]
+
+text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
+
+lemma zadd1_lemma:
+     "[| quorem((a,c),(aq,ar));  quorem((b,c),(bq,br));  c \<noteq> 0 |]  
+      ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))"
+by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib)
+
+(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
+lemma zdiv_zadd1_eq:
+     "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
+apply (case_tac "c = 0", simp)
+apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_div)
+done
+
+lemma zmod_zadd1_eq: "(a+b) mod (c::int) = (a mod c + b mod c) mod c"
+apply (case_tac "c = 0", simp)
+apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_mod)
+done
+
+lemma mod_div_trivial [simp]: "(a mod b) div b = (0::int)"
+apply (case_tac "b = 0", simp)
+apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial)
+done
+
+lemma mod_mod_trivial [simp]: "(a mod b) mod b = a mod (b::int)"
+apply (case_tac "b = 0", simp)
+apply (force simp add: linorder_neq_iff mod_pos_pos_trivial mod_neg_neg_trivial)
+done
+
+lemma zmod_zadd_left_eq: "(a+b) mod (c::int) = ((a mod c) + b) mod c"
+apply (rule trans [symmetric])
+apply (rule zmod_zadd1_eq, simp)
+apply (rule zmod_zadd1_eq [symmetric])
+done
+
+lemma zmod_zadd_right_eq: "(a+b) mod (c::int) = (a + (b mod c)) mod c"
+apply (rule trans [symmetric])
+apply (rule zmod_zadd1_eq, simp)
+apply (rule zmod_zadd1_eq [symmetric])
+done
+
+lemma zdiv_zadd_self1[simp]: "a \<noteq> (0::int) ==> (a+b) div a = b div a + 1"
+by (simp add: zdiv_zadd1_eq)
+
+lemma zdiv_zadd_self2[simp]: "a \<noteq> (0::int) ==> (b+a) div a = b div a + 1"
+by (simp add: zdiv_zadd1_eq)
+
+lemma zmod_zadd_self1[simp]: "(a+b) mod a = b mod (a::int)"
+apply (case_tac "a = 0", simp)
+apply (simp add: zmod_zadd1_eq)
+done
+
+lemma zmod_zadd_self2[simp]: "(b+a) mod a = b mod (a::int)"
+apply (case_tac "a = 0", simp)
+apply (simp add: zmod_zadd1_eq)
+done
+
+
+subsection{*Proving  @{term "a div (b*c) = (a div b) div c"} *}
+
+(*The condition c>0 seems necessary.  Consider that 7 div ~6 = ~2 but
+  7 div 2 div ~3 = 3 div ~3 = ~1.  The subcase (a div b) mod c = 0 seems
+  to cause particular problems.*)
+
+text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
+
+lemma zmult2_lemma_aux1: "[| (0::int) < c;  b < r;  r \<le> 0 |] ==> b*c < b*(q mod c) + r"
+apply (subgoal_tac "b * (c - q mod c) < r * 1")
+apply (simp add: right_diff_distrib)
+apply (rule order_le_less_trans)
+apply (erule_tac [2] mult_strict_right_mono)
+apply (rule mult_left_mono_neg)
+apply (auto simp add: compare_rls add_commute [of 1]
+                      add1_zle_eq pos_mod_bound)
+done
+
+lemma zmult2_lemma_aux2:
+     "[| (0::int) < c;   b < r;  r \<le> 0 |] ==> b * (q mod c) + r \<le> 0"
+apply (subgoal_tac "b * (q mod c) \<le> 0")
+ apply arith
+apply (simp add: mult_le_0_iff)
+done
+
+lemma zmult2_lemma_aux3: "[| (0::int) < c;  0 \<le> r;  r < b |] ==> 0 \<le> b * (q mod c) + r"
+apply (subgoal_tac "0 \<le> b * (q mod c) ")
+apply arith
+apply (simp add: zero_le_mult_iff)
+done
+
+lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c"
+apply (subgoal_tac "r * 1 < b * (c - q mod c) ")
+apply (simp add: right_diff_distrib)
+apply (rule order_less_le_trans)
+apply (erule mult_strict_right_mono)
+apply (rule_tac [2] mult_left_mono)
+apply (auto simp add: compare_rls add_commute [of 1]
+                      add1_zle_eq pos_mod_bound)
+done
+
+lemma zmult2_lemma: "[| quorem ((a,b), (q,r));  b \<noteq> 0;  0 < c |]  
+      ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))"
+by (auto simp add: mult_ac quorem_def linorder_neq_iff
+                   zero_less_mult_iff right_distrib [symmetric] 
+                   zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4)
+
+lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c"
+apply (case_tac "b = 0", simp)
+apply (force simp add: quorem_div_mod [THEN zmult2_lemma, THEN quorem_div])
+done
+
+lemma zmod_zmult2_eq:
+     "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"
+apply (case_tac "b = 0", simp)
+apply (force simp add: quorem_div_mod [THEN zmult2_lemma, THEN quorem_mod])
+done
+
+
+subsection{*Cancellation of Common Factors in div*}
+
+lemma zdiv_zmult_zmult1_aux1:
+     "[| (0::int) < b;  c \<noteq> 0 |] ==> (c*a) div (c*b) = a div b"
+by (subst zdiv_zmult2_eq, auto)
+
+lemma zdiv_zmult_zmult1_aux2:
+     "[| b < (0::int);  c \<noteq> 0 |] ==> (c*a) div (c*b) = a div b"
+apply (subgoal_tac " (c * (-a)) div (c * (-b)) = (-a) div (-b) ")
+apply (rule_tac [2] zdiv_zmult_zmult1_aux1, auto)
+done
+
+lemma zdiv_zmult_zmult1: "c \<noteq> (0::int) ==> (c*a) div (c*b) = a div b"
+apply (case_tac "b = 0", simp)
+apply (auto simp add: linorder_neq_iff zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2)
+done
+
+lemma zdiv_zmult_zmult2: "c \<noteq> (0::int) ==> (a*c) div (b*c) = a div b"
+apply (drule zdiv_zmult_zmult1)
+apply (auto simp add: mult_commute)
+done
+
+
+
+subsection{*Distribution of Factors over mod*}
+
+lemma zmod_zmult_zmult1_aux1:
+     "[| (0::int) < b;  c \<noteq> 0 |] ==> (c*a) mod (c*b) = c * (a mod b)"
+by (subst zmod_zmult2_eq, auto)
+
+lemma zmod_zmult_zmult1_aux2:
+     "[| b < (0::int);  c \<noteq> 0 |] ==> (c*a) mod (c*b) = c * (a mod b)"
+apply (subgoal_tac " (c * (-a)) mod (c * (-b)) = c * ((-a) mod (-b))")
+apply (rule_tac [2] zmod_zmult_zmult1_aux1, auto)
+done
+
+lemma zmod_zmult_zmult1: "(c*a) mod (c*b) = (c::int) * (a mod b)"
+apply (case_tac "b = 0", simp)
+apply (case_tac "c = 0", simp)
+apply (auto simp add: linorder_neq_iff zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2)
+done
+
+lemma zmod_zmult_zmult2: "(a*c) mod (b*c) = (a mod b) * (c::int)"
+apply (cut_tac c = c in zmod_zmult_zmult1)
+apply (auto simp add: mult_commute)
+done
+
+
+subsection {*Splitting Rules for div and mod*}
+
+text{*The proofs of the two lemmas below are essentially identical*}
+
+lemma split_pos_lemma:
+ "0<k ==> 
+    P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)"
+apply (rule iffI, clarify)
+ apply (erule_tac P="P ?x ?y" in rev_mp)  
+ apply (subst zmod_zadd1_eq) 
+ apply (subst zdiv_zadd1_eq) 
+ apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial)  
+txt{*converse direction*}
+apply (drule_tac x = "n div k" in spec) 
+apply (drule_tac x = "n mod k" in spec, simp)
+done
+
+lemma split_neg_lemma:
+ "k<0 ==>
+    P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)"
+apply (rule iffI, clarify)
+ apply (erule_tac P="P ?x ?y" in rev_mp)  
+ apply (subst zmod_zadd1_eq) 
+ apply (subst zdiv_zadd1_eq) 
+ apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial)  
+txt{*converse direction*}
+apply (drule_tac x = "n div k" in spec) 
+apply (drule_tac x = "n mod k" in spec, simp)
+done
+
+lemma split_zdiv:
+ "P(n div k :: int) =
+  ((k = 0 --> P 0) & 
+   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & 
+   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))"
+apply (case_tac "k=0", simp)
+apply (simp only: linorder_neq_iff)
+apply (erule disjE) 
+ apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] 
+                      split_neg_lemma [of concl: "%x y. P x"])
+done
+
+lemma split_zmod:
+ "P(n mod k :: int) =
+  ((k = 0 --> P n) & 
+   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & 
+   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))"
+apply (case_tac "k=0", simp)
+apply (simp only: linorder_neq_iff)
+apply (erule disjE) 
+ apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] 
+                      split_neg_lemma [of concl: "%x y. P y"])
+done
+
+(* Enable arith to deal with div 2 and mod 2: *)
+declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split]
+declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split]
+
+
+subsection{*Speeding up the Division Algorithm with Shifting*}
+
+text{*computing div by shifting *}
+
+lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a"
+proof cases
+  assume "a=0"
+    thus ?thesis by simp
+next
+  assume "a\<noteq>0" and le_a: "0\<le>a"   
+  hence a_pos: "1 \<le> a" by arith
+  hence one_less_a2: "1 < 2*a" by arith
+  hence le_2a: "2 * (1 + b mod a) \<le> 2 * a"
+    by (simp add: mult_le_cancel_left add_commute [of 1] add1_zle_eq)
+  with a_pos have "0 \<le> b mod a" by simp
+  hence le_addm: "0 \<le> 1 mod (2*a) + 2*(b mod a)"
+    by (simp add: mod_pos_pos_trivial one_less_a2)
+  with  le_2a
+  have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0"
+    by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2
+                  right_distrib) 
+  thus ?thesis
+    by (subst zdiv_zadd1_eq,
+        simp add: zdiv_zmult_zmult1 zmod_zmult_zmult1 one_less_a2
+                  div_pos_pos_trivial)
+qed
+
+lemma neg_zdiv_mult_2: "a \<le> (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a"
+apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ")
+apply (rule_tac [2] pos_zdiv_mult_2)
+apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib)
+apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))")
+apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric],
+       simp) 
+done
+
+
+(*Not clear why this must be proved separately; probably number_of causes
+  simplification problems*)
+lemma not_0_le_lemma: "~ 0 \<le> x ==> x \<le> (0::int)"
+by auto
+
+lemma zdiv_number_of_BIT[simp]:
+     "number_of (v BIT b) div number_of (w BIT bit.B0) =  
+          (if b=bit.B0 | (0::int) \<le> number_of w                    
+           then number_of v div (number_of w)     
+           else (number_of v + (1::int)) div (number_of w))"
+apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) 
+apply (simp add: zdiv_zmult_zmult1 pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac 
+            split: bit.split)
+done
+
+
+subsection{*Computing mod by Shifting (proofs resemble those for div)*}
+
+lemma pos_zmod_mult_2:
+     "(0::int) \<le> a ==> (1 + 2*b) mod (2*a) = 1 + 2 * (b mod a)"
+apply (case_tac "a = 0", simp)
+apply (subgoal_tac "1 < a * 2")
+ prefer 2 apply arith
+apply (subgoal_tac "2* (1 + b mod a) \<le> 2*a")
+ apply (rule_tac [2] mult_left_mono)
+apply (auto simp add: add_commute [of 1] mult_commute add1_zle_eq 
+                      pos_mod_bound)
+apply (subst zmod_zadd1_eq)
+apply (simp add: zmod_zmult_zmult2 mod_pos_pos_trivial)
+apply (rule mod_pos_pos_trivial)
+apply (auto simp add: mod_pos_pos_trivial left_distrib)
+apply (subgoal_tac "0 \<le> b mod a", arith, simp)
+done
+
+lemma neg_zmod_mult_2:
+     "a \<le> (0::int) ==> (1 + 2*b) mod (2*a) = 2 * ((b+1) mod a) - 1"
+apply (subgoal_tac "(1 + 2* (-b - 1)) mod (2* (-a)) = 
+                    1 + 2* ((-b - 1) mod (-a))")
+apply (rule_tac [2] pos_zmod_mult_2)
+apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib)
+apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))")
+ prefer 2 apply simp 
+apply (simp only: zmod_zminus_zminus diff_minus minus_add_distrib [symmetric])
+done
+
+lemma zmod_number_of_BIT [simp]:
+     "number_of (v BIT b) mod number_of (w BIT bit.B0) =  
+      (case b of
+          bit.B0 => 2 * (number_of v mod number_of w)
+        | bit.B1 => if (0::int) \<le> number_of w  
+                then 2 * (number_of v mod number_of w) + 1     
+                else 2 * ((number_of v + (1::int)) mod number_of w) - 1)"
+apply (simp only: number_of_eq numeral_simps UNIV_I split: bit.split) 
+apply (simp add: zmod_zmult_zmult1 pos_zmod_mult_2 
+                 not_0_le_lemma neg_zmod_mult_2 add_ac)
+done
+
+
+subsection{*Quotients of Signs*}
+
+lemma div_neg_pos_less0: "[| a < (0::int);  0 < b |] ==> a div b < 0"
+apply (subgoal_tac "a div b \<le> -1", force)
+apply (rule order_trans)
+apply (rule_tac a' = "-1" in zdiv_mono1)
+apply (auto simp add: zdiv_minus1)
+done
+
+lemma div_nonneg_neg_le0: "[| (0::int) \<le> a;  b < 0 |] ==> a div b \<le> 0"
+by (drule zdiv_mono1_neg, auto)
+
+lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)"
+apply auto
+apply (drule_tac [2] zdiv_mono1)
+apply (auto simp add: linorder_neq_iff)
+apply (simp (no_asm_use) add: linorder_not_less [symmetric])
+apply (blast intro: div_neg_pos_less0)
+done
+
+lemma neg_imp_zdiv_nonneg_iff:
+     "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))"
+apply (subst zdiv_zminus_zminus [symmetric])
+apply (subst pos_imp_zdiv_nonneg_iff, auto)
+done
+
+(*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*)
+lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)"
+by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff)
+
+(*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*)
+lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)"
+by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff)
+
+
+subsection {* The Divides Relation *}
+
+lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))"
+by(simp add:dvd_def zmod_eq_0_iff)
+
+lemmas zdvd_iff_zmod_eq_0_number_of [simp] =
+  zdvd_iff_zmod_eq_0 [of "number_of x" "number_of y", standard]
+
+lemma zdvd_0_right [iff]: "(m::int) dvd 0"
+by (simp add: dvd_def)
+
+lemma zdvd_0_left [iff]: "(0 dvd (m::int)) = (m = 0)"
+  by (simp add: dvd_def)
+
+lemma zdvd_1_left [iff]: "1 dvd (m::int)"
+  by (simp add: dvd_def)
+
+lemma zdvd_refl [simp]: "m dvd (m::int)"
+by (auto simp add: dvd_def intro: zmult_1_right [symmetric])
+
+lemma zdvd_trans: "m dvd n ==> n dvd k ==> m dvd (k::int)"
+by (auto simp add: dvd_def intro: mult_assoc)
+
+lemma zdvd_zminus_iff: "(m dvd -n) = (m dvd (n::int))"
+  apply (simp add: dvd_def, auto)
+   apply (rule_tac [!] x = "-k" in exI, auto)
+  done
+
+lemma zdvd_zminus2_iff: "(-m dvd n) = (m dvd (n::int))"
+  apply (simp add: dvd_def, auto)
+   apply (rule_tac [!] x = "-k" in exI, auto)
+  done
+lemma zdvd_abs1: "( \<bar>i::int\<bar> dvd j) = (i dvd j)" 
+  apply (cases "i > 0", simp)
+  apply (simp add: dvd_def)
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x="- k" in exI, simp)
+  apply (erule exE)
+  apply (rule_tac x="- k" in exI, simp)
+  done
+lemma zdvd_abs2: "( (i::int) dvd \<bar>j\<bar>) = (i dvd j)" 
+  apply (cases "j > 0", simp)
+  apply (simp add: dvd_def)
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x="- k" in exI, simp)
+  apply (erule exE)
+  apply (rule_tac x="- k" in exI, simp)
+  done
+
+lemma zdvd_anti_sym:
+    "0 < m ==> 0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)"
+  apply (simp add: dvd_def, auto)
+  apply (simp add: mult_assoc zero_less_mult_iff zmult_eq_1_iff)
+  done
+
+lemma zdvd_zadd: "k dvd m ==> k dvd n ==> k dvd (m + n :: int)"
+  apply (simp add: dvd_def)
+  apply (blast intro: right_distrib [symmetric])
+  done
+
+lemma zdvd_dvd_eq: assumes anz:"a \<noteq> 0" and ab: "(a::int) dvd b" and ba:"b dvd a" 
+  shows "\<bar>a\<bar> = \<bar>b\<bar>"
+proof-
+  from ab obtain k where k:"b = a*k" unfolding dvd_def by blast 
+  from ba obtain k' where k':"a = b*k'" unfolding dvd_def by blast 
+  from k k' have "a = a*k*k'" by simp
+  with mult_cancel_left1[where c="a" and b="k*k'"]
+  have kk':"k*k' = 1" using anz by (simp add: mult_assoc)
+  hence "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1" by (simp add: zmult_eq_1_iff)
+  thus ?thesis using k k' by auto
+qed
+
+lemma zdvd_zdiff: "k dvd m ==> k dvd n ==> k dvd (m - n :: int)"
+  apply (simp add: dvd_def)
+  apply (blast intro: right_diff_distrib [symmetric])
+  done
+
+lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)"
+  apply (subgoal_tac "m = n + (m - n)")
+   apply (erule ssubst)
+   apply (blast intro: zdvd_zadd, simp)
+  done
+
+lemma zdvd_zmult: "k dvd (n::int) ==> k dvd m * n"
+  apply (simp add: dvd_def)
+  apply (blast intro: mult_left_commute)
+  done
+
+lemma zdvd_zmult2: "k dvd (m::int) ==> k dvd m * n"
+  apply (subst mult_commute)
+  apply (erule zdvd_zmult)
+  done
+
+lemma zdvd_triv_right [iff]: "(k::int) dvd m * k"
+  apply (rule zdvd_zmult)
+  apply (rule zdvd_refl)
+  done
+
+lemma zdvd_triv_left [iff]: "(k::int) dvd k * m"
+  apply (rule zdvd_zmult2)
+  apply (rule zdvd_refl)
+  done
+
+lemma zdvd_zmultD2: "j * k dvd n ==> j dvd (n::int)"
+  apply (simp add: dvd_def)
+  apply (simp add: mult_assoc, blast)
+  done
+
+lemma zdvd_zmultD: "j * k dvd n ==> k dvd (n::int)"
+  apply (rule zdvd_zmultD2)
+  apply (subst mult_commute, assumption)
+  done
+
+lemma zdvd_zmult_mono: "i dvd m ==> j dvd (n::int) ==> i * j dvd m * n"
+  apply (simp add: dvd_def, clarify)
+  apply (rule_tac x = "k * ka" in exI)
+  apply (simp add: mult_ac)
+  done
+
+lemma zdvd_reduce: "(k dvd n + k * m) = (k dvd (n::int))"
+  apply (rule iffI)
+   apply (erule_tac [2] zdvd_zadd)
+   apply (subgoal_tac "n = (n + k * m) - k * m")
+    apply (erule ssubst)
+    apply (erule zdvd_zdiff, simp_all)
+  done
+
+lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n"
+  apply (simp add: dvd_def)
+  apply (auto simp add: zmod_zmult_zmult1)
+  done
+
+lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)"
+  apply (subgoal_tac "k dvd n * (m div n) + m mod n")
+   apply (simp add: zmod_zdiv_equality [symmetric])
+  apply (simp only: zdvd_zadd zdvd_zmult2)
+  done
+
+lemma zdvd_not_zless: "0 < m ==> m < n ==> \<not> n dvd (m::int)"
+  apply (simp add: dvd_def, auto)
+  apply (subgoal_tac "0 < n")
+   prefer 2
+   apply (blast intro: order_less_trans)
+  apply (simp add: zero_less_mult_iff)
+  apply (subgoal_tac "n * k < n * 1")
+   apply (drule mult_less_cancel_left [THEN iffD1], auto)
+  done
+lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)"
+  using zmod_zdiv_equality[where a="m" and b="n"]
+  by (simp add: ring_eq_simps)
+
+lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m"
+apply (subgoal_tac "m mod n = 0")
+ apply (simp add: zmult_div_cancel)
+apply (simp only: zdvd_iff_zmod_eq_0)
+done
+
+lemma zdvd_mult_cancel: assumes d:"k * m dvd k * n" and kz:"k \<noteq> (0::int)"
+  shows "m dvd n"
+proof-
+  from d obtain h where h: "k*n = k*m * h" unfolding dvd_def by blast
+  {assume "n \<noteq> m*h" hence "k* n \<noteq> k* (m*h)" using kz by simp
+    with h have False by (simp add: mult_assoc)}
+  hence "n = m * h" by blast
+  thus ?thesis by blast
+qed
+
+theorem ex_nat: "(\<exists>x::nat. P x) = (\<exists>x::int. 0 <= x \<and> P (nat x))"
+  apply (simp split add: split_nat)
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x = "int x" in exI)
+  apply simp
+  apply (erule exE)
+  apply (rule_tac x = "nat x" in exI)
+  apply (erule conjE)
+  apply (erule_tac x = "nat x" in allE)
+  apply simp
+  done
+
+theorem zdvd_int: "(x dvd y) = (int x dvd int y)"
+  apply (simp only: dvd_def ex_nat int_int_eq [symmetric] zmult_int [symmetric]
+    nat_0_le cong add: conj_cong)
+  apply (rule iffI)
+  apply iprover
+  apply (erule exE)
+  apply (case_tac "x=0")
+  apply (rule_tac x=0 in exI)
+  apply simp
+  apply (case_tac "0 \<le> k")
+  apply iprover
+  apply (simp add: linorder_not_le)
+  apply (drule mult_strict_left_mono_neg [OF iffD2 [OF zero_less_int_conv]])
+  apply assumption
+  apply (simp add: mult_ac)
+  done
+
+lemma zdvd1_eq[simp]: "(x::int) dvd 1 = ( \<bar>x\<bar> = 1)"
+proof
+  assume d: "x dvd 1" hence "int (nat \<bar>x\<bar>) dvd int (nat 1)" by (simp add: zdvd_abs1)
+  hence "nat \<bar>x\<bar> dvd 1" by (simp add: zdvd_int)
+  hence "nat \<bar>x\<bar> = 1"  by simp
+  thus "\<bar>x\<bar> = 1" by (cases "x < 0", auto)
+next
+  assume "\<bar>x\<bar>=1" thus "x dvd 1" 
+    by(cases "x < 0",simp_all add: minus_equation_iff zdvd_iff_zmod_eq_0)
+qed
+lemma zdvd_mult_cancel1: 
+  assumes mp:"m \<noteq>(0::int)" shows "(m * n dvd m) = (\<bar>n\<bar> = 1)"
+proof
+  assume n1: "\<bar>n\<bar> = 1" thus "m * n dvd m" 
+    by (cases "n >0", auto simp add: zdvd_zminus2_iff minus_equation_iff)
+next
+  assume H: "m * n dvd m" hence H2: "m * n dvd m * 1" by simp
+  from zdvd_mult_cancel[OF H2 mp] show "\<bar>n\<bar> = 1" by (simp only: zdvd1_eq)
+qed
+
+lemma int_dvd_iff: "(int m dvd z) = (m dvd nat (abs z))"
+  apply (auto simp add: dvd_def nat_abs_mult_distrib)
+  apply (auto simp add: nat_eq_iff abs_if split add: split_if_asm)
+   apply (rule_tac x = "-(int k)" in exI)
+  apply (auto simp add: int_mult)
+  done
+
+lemma dvd_int_iff: "(z dvd int m) = (nat (abs z) dvd m)"
+  apply (auto simp add: dvd_def abs_if int_mult)
+    apply (rule_tac [3] x = "nat k" in exI)
+    apply (rule_tac [2] x = "-(int k)" in exI)
+    apply (rule_tac x = "nat (-k)" in exI)
+    apply (cut_tac [3] k = m in int_less_0_conv)
+    apply (cut_tac k = m in int_less_0_conv)
+    apply (auto simp add: zero_le_mult_iff mult_less_0_iff
+      nat_mult_distrib [symmetric] nat_eq_iff2)
+  done
+
+lemma nat_dvd_iff: "(nat z dvd m) = (if 0 \<le> z then (z dvd int m) else m = 0)"
+  apply (auto simp add: dvd_def int_mult)
+  apply (rule_tac x = "nat k" in exI)
+  apply (cut_tac k = m in int_less_0_conv)
+  apply (auto simp add: zero_le_mult_iff mult_less_0_iff
+    nat_mult_distrib [symmetric] nat_eq_iff2)
+  done
+
+lemma zminus_dvd_iff [iff]: "(-z dvd w) = (z dvd (w::int))"
+  apply (auto simp add: dvd_def)
+   apply (rule_tac [!] x = "-k" in exI, auto)
+  done
+
+lemma dvd_zminus_iff [iff]: "(z dvd -w) = (z dvd (w::int))"
+  apply (auto simp add: dvd_def)
+   apply (drule minus_equation_iff [THEN iffD1])
+   apply (rule_tac [!] x = "-k" in exI, auto)
+  done
+
+lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \<le> (n::int)"
+  apply (rule_tac z=n in int_cases)
+  apply (auto simp add: dvd_int_iff) 
+  apply (rule_tac z=z in int_cases) 
+  apply (auto simp add: dvd_imp_le) 
+  done
+
+
+subsection{*Integer Powers*} 
+
+instance int :: power ..
+
+primrec
+  "p ^ 0 = 1"
+  "p ^ (Suc n) = (p::int) * (p ^ n)"
+
+
+instance int :: recpower
+proof
+  fix z :: int
+  fix n :: nat
+  show "z^0 = 1" by simp
+  show "z^(Suc n) = z * (z^n)" by simp
+qed
+
+
+lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m"
+apply (induct "y", auto)
+apply (rule zmod_zmult1_eq [THEN trans])
+apply (simp (no_asm_simp))
+apply (rule zmod_zmult_distrib [symmetric])
+done
+
+lemma zpower_zadd_distrib: "x^(y+z) = ((x^y)*(x^z)::int)"
+  by (rule Power.power_add)
+
+lemma zpower_zpower: "(x^y)^z = (x^(y*z)::int)"
+  by (rule Power.power_mult [symmetric])
+
+lemma zero_less_zpower_abs_iff [simp]:
+     "(0 < (abs x)^n) = (x \<noteq> (0::int) | n=0)"
+apply (induct "n")
+apply (auto simp add: zero_less_mult_iff)
+done
+
+lemma zero_le_zpower_abs [simp]: "(0::int) <= (abs x)^n"
+apply (induct "n")
+apply (auto simp add: zero_le_mult_iff)
+done
+
+lemma int_power: "int (m^n) = (int m) ^ n"
+  by (induct n, simp_all add: int_mult)
+
+text{*Compatibility binding*}
+lemmas zpower_int = int_power [symmetric]
+
+lemma zdiv_int: "int (a div b) = (int a) div (int b)"
+apply (subst split_div, auto)
+apply (subst split_zdiv, auto)
+apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient)
+apply (auto simp add: IntDiv.quorem_def int_eq_of_nat)
+done
+
+lemma zmod_int: "int (a mod b) = (int a) mod (int b)"
+apply (subst split_mod, auto)
+apply (subst split_zmod, auto)
+apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia 
+       in unique_remainder)
+apply (auto simp add: IntDiv.quorem_def int_eq_of_nat)
+done
+
+text{*Suggested by Matthias Daum*}
+lemma int_power_div_base:
+     "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"
+apply (subgoal_tac "k ^ m = k ^ ((m - 1) + 1)")
+ apply (erule ssubst)
+ apply (simp only: power_add)
+ apply simp_all
+done
+
+text {* code serializer setup *}
+
+code_modulename SML
+  IntDiv Integer
+
+code_modulename OCaml
+  IntDiv Integer
+
+code_modulename Haskell
+  IntDiv Divides
+
+end
--- a/src/HOL/Integ/IntArith.thy	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,411 +0,0 @@
-(*  Title:      HOL/Integ/IntArith.thy
-    ID:         $Id$
-    Authors:    Larry Paulson and Tobias Nipkow
-*)
-
-header {* Integer arithmetic *}
-
-theory IntArith
-imports Numeral "../Wellfounded_Relations"
-uses "~~/src/Provers/Arith/assoc_fold.ML" ("int_arith1.ML")
-begin
-
-text{*Duplicate: can't understand why it's necessary*}
-declare numeral_0_eq_0 [simp]
-
-
-subsection{*Inequality Reasoning for the Arithmetic Simproc*}
-
-lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)"
-by simp 
-
-lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)"
-by simp
-
-lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)"
-by simp 
-
-lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)"
-by simp
-
-lemma divide_numeral_1: "a / Numeral1 = (a::'a::{number_ring,field})"
-by simp
-
-lemma inverse_numeral_1:
-  "inverse Numeral1 = (Numeral1::'a::{number_ring,field})"
-by simp
-
-text{*Theorem lists for the cancellation simprocs. The use of binary numerals
-for 0 and 1 reduces the number of special cases.*}
-
-lemmas add_0s = add_numeral_0 add_numeral_0_right
-lemmas mult_1s = mult_numeral_1 mult_numeral_1_right 
-                 mult_minus1 mult_minus1_right
-
-
-subsection{*Special Arithmetic Rules for Abstract 0 and 1*}
-
-text{*Arithmetic computations are defined for binary literals, which leaves 0
-and 1 as special cases. Addition already has rules for 0, but not 1.
-Multiplication and unary minus already have rules for both 0 and 1.*}
-
-
-lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'"
-by simp
-
-
-lemmas add_number_of_eq = number_of_add [symmetric]
-
-text{*Allow 1 on either or both sides*}
-lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)"
-by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq)
-
-lemmas add_special =
-    one_add_one_is_two
-    binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard]
-    binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard]
-
-text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*}
-lemmas diff_special =
-    binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard]
-    binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard]
-
-text{*Allow 0 or 1 on either side with a binary numeral on the other*}
-lemmas eq_special =
-    binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard]
-    binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard]
-    binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard]
-    binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard]
-
-text{*Allow 0 or 1 on either side with a binary numeral on the other*}
-lemmas less_special =
-  binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard]
-  binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard]
-  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard]
-  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard]
-
-text{*Allow 0 or 1 on either side with a binary numeral on the other*}
-lemmas le_special =
-    binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard]
-    binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard]
-    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard]
-    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard]
-
-lemmas arith_special[simp] = 
-       add_special diff_special eq_special less_special le_special
-
-
-lemma min_max_01: "min (0::int) 1 = 0 & min (1::int) 0 = 0 &
-                   max (0::int) 1 = 1 & max (1::int) 0 = 1"
-by(simp add:min_def max_def)
-
-lemmas min_max_special[simp] =
- min_max_01
- max_def[of "0::int" "number_of v", standard, simp]
- min_def[of "0::int" "number_of v", standard, simp]
- max_def[of "number_of u" "0::int", standard, simp]
- min_def[of "number_of u" "0::int", standard, simp]
- max_def[of "1::int" "number_of v", standard, simp]
- min_def[of "1::int" "number_of v", standard, simp]
- max_def[of "number_of u" "1::int", standard, simp]
- min_def[of "number_of u" "1::int", standard, simp]
-
-use "int_arith1.ML"
-setup int_arith_setup
-
-
-subsection{*Lemmas About Small Numerals*}
-
-lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)"
-proof -
-  have "(of_int -1 :: 'a) = of_int (- 1)" by simp
-  also have "... = - of_int 1" by (simp only: of_int_minus)
-  also have "... = -1" by simp
-  finally show ?thesis .
-qed
-
-lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_idom,number_ring})"
-by (simp add: abs_if)
-
-lemma abs_power_minus_one [simp]:
-     "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,recpower})"
-by (simp add: power_abs)
-
-lemma of_int_number_of_eq:
-     "of_int (number_of v) = (number_of v :: 'a :: number_ring)"
-by (simp add: number_of_eq) 
-
-text{*Lemmas for specialist use, NOT as default simprules*}
-lemma mult_2: "2 * z = (z+z::'a::number_ring)"
-proof -
-  have "2*z = (1 + 1)*z" by simp
-  also have "... = z+z" by (simp add: left_distrib)
-  finally show ?thesis .
-qed
-
-lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)"
-by (subst mult_commute, rule mult_2)
-
-
-subsection{*More Inequality Reasoning*}
-
-lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)"
-by arith
-
-lemma add1_zle_eq: "(w + (1::int) \<le> z) = (w<z)"
-by arith
-
-lemma zle_diff1_eq [simp]: "(w \<le> z - (1::int)) = (w<z)"
-by arith
-
-lemma zle_add1_eq_le [simp]: "(w < z + (1::int)) = (w\<le>z)"
-by arith
-
-lemma int_one_le_iff_zero_less: "((1::int) \<le> z) = (0 < z)"
-by arith
-
-
-subsection{*The Functions @{term nat} and @{term int}*}
-
-text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and
-  @{term "w + - z"}*}
-declare Zero_int_def [symmetric, simp]
-declare One_int_def [symmetric, simp]
-
-lemmas diff_int_def_symmetric = diff_int_def [symmetric, simp]
-
-lemma nat_0: "nat 0 = 0"
-by (simp add: nat_eq_iff)
-
-lemma nat_1: "nat 1 = Suc 0"
-by (subst nat_eq_iff, simp)
-
-lemma nat_2: "nat 2 = Suc (Suc 0)"
-by (subst nat_eq_iff, simp)
-
-lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
-apply (insert zless_nat_conj [of 1 z])
-apply (auto simp add: nat_1)
-done
-
-text{*This simplifies expressions of the form @{term "int n = z"} where
-      z is an integer literal.*}
-lemmas int_eq_iff_number_of [simp] = int_eq_iff [of _ "number_of v", standard]
-
-
-lemma split_nat [arith_split]:
-  "P(nat(i::int)) = ((\<forall>n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))"
-  (is "?P = (?L & ?R)")
-proof (cases "i < 0")
-  case True thus ?thesis by simp
-next
-  case False
-  have "?P = ?L"
-  proof
-    assume ?P thus ?L using False by clarsimp
-  next
-    assume ?L thus ?P using False by simp
-  qed
-  with False show ?thesis by simp
-qed
-
-
-(*Analogous to zadd_int*)
-lemma zdiff_int: "n \<le> m ==> int m - int n = int (m-n)" 
-by (induct m n rule: diff_induct, simp_all)
-
-lemma nat_mult_distrib: "(0::int) \<le> z ==> nat (z*z') = nat z * nat z'"
-apply (cases "0 \<le> z'")
-apply (rule inj_int [THEN injD])
-apply (simp add: int_mult zero_le_mult_iff)
-apply (simp add: mult_le_0_iff)
-done
-
-lemma nat_mult_distrib_neg: "z \<le> (0::int) ==> nat(z*z') = nat(-z) * nat(-z')"
-apply (rule trans)
-apply (rule_tac [2] nat_mult_distrib, auto)
-done
-
-lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)"
-apply (cases "z=0 | w=0")
-apply (auto simp add: abs_if nat_mult_distrib [symmetric] 
-                      nat_mult_distrib_neg [symmetric] mult_less_0_iff)
-done
-
-
-subsection "Induction principles for int"
-
-text{*Well-founded segments of the integers*}
-
-definition
-  int_ge_less_than  ::  "int => (int * int) set"
-where
-  "int_ge_less_than d = {(z',z). d \<le> z' & z' < z}"
-
-theorem wf_int_ge_less_than: "wf (int_ge_less_than d)"
-proof -
-  have "int_ge_less_than d \<subseteq> measure (%z. nat (z-d))"
-    by (auto simp add: int_ge_less_than_def)
-  thus ?thesis 
-    by (rule wf_subset [OF wf_measure]) 
-qed
-
-text{*This variant looks odd, but is typical of the relations suggested
-by RankFinder.*}
-
-definition
-  int_ge_less_than2 ::  "int => (int * int) set"
-where
-  "int_ge_less_than2 d = {(z',z). d \<le> z & z' < z}"
-
-theorem wf_int_ge_less_than2: "wf (int_ge_less_than2 d)"
-proof -
-  have "int_ge_less_than2 d \<subseteq> measure (%z. nat (1+z-d))" 
-    by (auto simp add: int_ge_less_than2_def)
-  thus ?thesis 
-    by (rule wf_subset [OF wf_measure]) 
-qed
-
-                     (* `set:int': dummy construction *)
-theorem int_ge_induct[case_names base step,induct set:int]:
-  assumes ge: "k \<le> (i::int)" and
-        base: "P(k)" and
-        step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
-  shows "P i"
-proof -
-  { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i"
-    proof (induct n)
-      case 0
-      hence "i = k" by arith
-      thus "P i" using base by simp
-    next
-      case (Suc n)
-      hence "n = nat((i - 1) - k)" by arith
-      moreover
-      have ki1: "k \<le> i - 1" using Suc.prems by arith
-      ultimately
-      have "P(i - 1)" by(rule Suc.hyps)
-      from step[OF ki1 this] show ?case by simp
-    qed
-  }
-  with ge show ?thesis by fast
-qed
-
-                     (* `set:int': dummy construction *)
-theorem int_gr_induct[case_names base step,induct set:int]:
-  assumes gr: "k < (i::int)" and
-        base: "P(k+1)" and
-        step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
-  shows "P i"
-apply(rule int_ge_induct[of "k + 1"])
-  using gr apply arith
- apply(rule base)
-apply (rule step, simp+)
-done
-
-theorem int_le_induct[consumes 1,case_names base step]:
-  assumes le: "i \<le> (k::int)" and
-        base: "P(k)" and
-        step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
-  shows "P i"
-proof -
-  { fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i"
-    proof (induct n)
-      case 0
-      hence "i = k" by arith
-      thus "P i" using base by simp
-    next
-      case (Suc n)
-      hence "n = nat(k - (i+1))" by arith
-      moreover
-      have ki1: "i + 1 \<le> k" using Suc.prems by arith
-      ultimately
-      have "P(i+1)" by(rule Suc.hyps)
-      from step[OF ki1 this] show ?case by simp
-    qed
-  }
-  with le show ?thesis by fast
-qed
-
-theorem int_less_induct [consumes 1,case_names base step]:
-  assumes less: "(i::int) < k" and
-        base: "P(k - 1)" and
-        step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
-  shows "P i"
-apply(rule int_le_induct[of _ "k - 1"])
-  using less apply arith
- apply(rule base)
-apply (rule step, simp+)
-done
-
-subsection{*Intermediate value theorems*}
-
-lemma int_val_lemma:
-     "(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) -->  
-      f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))"
-apply (induct_tac "n", simp)
-apply (intro strip)
-apply (erule impE, simp)
-apply (erule_tac x = n in allE, simp)
-apply (case_tac "k = f (n+1) ")
- apply force
-apply (erule impE)
- apply (simp add: abs_if split add: split_if_asm)
-apply (blast intro: le_SucI)
-done
-
-lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)]
-
-lemma nat_intermed_int_val:
-     "[| \<forall>i. m \<le> i & i < n --> abs(f(i + 1::nat) - f i) \<le> 1; m < n;  
-         f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)"
-apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k 
-       in int_val_lemma)
-apply simp
-apply (erule exE)
-apply (rule_tac x = "i+m" in exI, arith)
-done
-
-
-subsection{*Products and 1, by T. M. Rasmussen*}
-
-lemma zabs_less_one_iff [simp]: "(\<bar>z\<bar> < 1) = (z = (0::int))"
-by arith
-
-lemma abs_zmult_eq_1: "(\<bar>m * n\<bar> = 1) ==> \<bar>m\<bar> = (1::int)"
-apply (cases "\<bar>n\<bar>=1") 
-apply (simp add: abs_mult) 
-apply (rule ccontr) 
-apply (auto simp add: linorder_neq_iff abs_mult) 
-apply (subgoal_tac "2 \<le> \<bar>m\<bar> & 2 \<le> \<bar>n\<bar>")
- prefer 2 apply arith 
-apply (subgoal_tac "2*2 \<le> \<bar>m\<bar> * \<bar>n\<bar>", simp) 
-apply (rule mult_mono, auto) 
-done
-
-lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1"
-by (insert abs_zmult_eq_1 [of m n], arith)
-
-lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)"
-apply (auto dest: pos_zmult_eq_1_iff_lemma) 
-apply (simp add: mult_commute [of m]) 
-apply (frule pos_zmult_eq_1_iff_lemma, auto) 
-done
-
-lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))"
-apply (rule iffI) 
- apply (frule pos_zmult_eq_1_iff_lemma)
- apply (simp add: mult_commute [of m]) 
- apply (frule pos_zmult_eq_1_iff_lemma, auto) 
-done
-
-
-subsection {* Legacy ML bindings *}
-
-ML {*
-val of_int_number_of_eq = @{thm of_int_number_of_eq};
-val nat_0 = @{thm nat_0};
-val nat_1 = @{thm nat_1};
-*}
-
-end
--- a/src/HOL/Integ/IntDef.thy	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,890 +0,0 @@
-(*  Title:      IntDef.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1996  University of Cambridge
-
-*)
-
-header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} 
-
-theory IntDef
-imports Equiv_Relations Nat
-begin
-
-text {* the equivalence relation underlying the integers *}
-
-definition
-  intrel :: "((nat \<times> nat) \<times> (nat \<times> nat)) set"
-where
-  "intrel = {((x, y), (u, v)) | x y u v. x + v = u +y }"
-
-typedef (Integ)
-  int = "UNIV//intrel"
-  by (auto simp add: quotient_def)
-
-definition
-  int :: "nat \<Rightarrow> int"
-where
-  [code func del]: "int m = Abs_Integ (intrel `` {(m, 0)})"
-
-instance int :: zero
-  Zero_int_def: "0 \<equiv> int 0" ..
-
-instance int :: one
-  One_int_def: "1 \<equiv> int 1" ..
-
-instance int :: plus
-  add_int_def: "z + w \<equiv> Abs_Integ
-    (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u, v) \<in> Rep_Integ w.
-      intrel `` {(x + u, y + v)})" ..
-
-instance int :: minus
-  minus_int_def:
-    "- z \<equiv> Abs_Integ (\<Union>(x, y) \<in> Rep_Integ z. intrel `` {(y, x)})"
-  diff_int_def:  "z - w \<equiv> z + (-w)" ..
-
-instance int :: times
-  mult_int_def: "z * w \<equiv>  Abs_Integ
-    (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u,v ) \<in> Rep_Integ w.
-      intrel `` {(x*u + y*v, x*v + y*u)})" ..
-
-instance int :: ord
-  le_int_def:
-   "z \<le> w \<equiv> \<exists>x y u v. x+v \<le> u+y \<and> (x, y) \<in> Rep_Integ z \<and> (u, v) \<in> Rep_Integ w"
-  less_int_def: "z < w \<equiv> z \<le> w \<and> z \<noteq> w" ..
-
-lemmas [code func del] = Zero_int_def One_int_def add_int_def
-  minus_int_def mult_int_def le_int_def less_int_def
-
-
-subsection{*Construction of the Integers*}
-
-subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
-
-lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)"
-by (simp add: intrel_def)
-
-lemma equiv_intrel: "equiv UNIV intrel"
-by (simp add: intrel_def equiv_def refl_def sym_def trans_def)
-
-text{*Reduces equality of equivalence classes to the @{term intrel} relation:
-  @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
-lemmas equiv_intrel_iff [simp] = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I]
-
-text{*All equivalence classes belong to set of representatives*}
-lemma [simp]: "intrel``{(x,y)} \<in> Integ"
-by (auto simp add: Integ_def intrel_def quotient_def)
-
-text{*Reduces equality on abstractions to equality on representatives:
-  @{prop "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
-declare Abs_Integ_inject [simp]  Abs_Integ_inverse [simp]
-
-text{*Case analysis on the representation of an integer as an equivalence
-      class of pairs of naturals.*}
-lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]:
-     "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
-apply (rule Abs_Integ_cases [of z]) 
-apply (auto simp add: Integ_def quotient_def) 
-done
-
-
-subsubsection{*@{term int}: Embedding the Naturals into the Integers*}
-
-lemma inj_int: "inj int"
-by (simp add: inj_on_def int_def)
-
-lemma int_int_eq [iff]: "(int m = int n) = (m = n)"
-by (fast elim!: inj_int [THEN injD])
-
-
-subsubsection{*Integer Unary Negation*}
-
-lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
-proof -
-  have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
-    by (simp add: congruent_def) 
-  thus ?thesis
-    by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel])
-qed
-
-lemma zminus_zminus: "- (- z) = (z::int)"
-  by (cases z) (simp add: minus)
-
-lemma zminus_0: "- 0 = (0::int)"
-  by (simp add: int_def Zero_int_def minus)
-
-
-subsection{*Integer Addition*}
-
-lemma add:
-     "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
-      Abs_Integ (intrel``{(x+u, y+v)})"
-proof -
-  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z) 
-        respects2 intrel"
-    by (simp add: congruent2_def)
-  thus ?thesis
-    by (simp add: add_int_def UN_UN_split_split_eq
-                  UN_equiv_class2 [OF equiv_intrel equiv_intrel])
-qed
-
-lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)"
-  by (cases z, cases w) (simp add: minus add)
-
-lemma zadd_commute: "(z::int) + w = w + z"
-  by (cases z, cases w) (simp add: add_ac add)
-
-lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
-  by (cases z1, cases z2, cases z3) (simp add: add add_assoc)
-
-(*For AC rewriting*)
-lemma zadd_left_commute: "x + (y + z) = y + ((x + z)  ::int)"
-  apply (rule mk_left_commute [of "op +"])
-  apply (rule zadd_assoc)
-  apply (rule zadd_commute)
-  done
-
-lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute
-
-lemmas zmult_ac = OrderedGroup.mult_ac
-
-lemma zadd_int: "(int m) + (int n) = int (m + n)"
-  by (simp add: int_def add)
-
-lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z"
-  by (simp add: zadd_int zadd_assoc [symmetric])
-
-(*also for the instance declaration int :: comm_monoid_add*)
-lemma zadd_0: "(0::int) + z = z"
-apply (simp add: Zero_int_def int_def)
-apply (cases z, simp add: add)
-done
-
-lemma zadd_0_right: "z + (0::int) = z"
-by (rule trans [OF zadd_commute zadd_0])
-
-lemma zadd_zminus_inverse2: "(- z) + z = (0::int)"
-by (cases z, simp add: int_def Zero_int_def minus add)
-
-
-subsection{*Integer Multiplication*}
-
-text{*Congruence property for multiplication*}
-lemma mult_congruent2:
-     "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)
-      respects2 intrel"
-apply (rule equiv_intrel [THEN congruent2_commuteI])
- apply (force simp add: mult_ac, clarify) 
-apply (simp add: congruent_def mult_ac)  
-apply (rename_tac u v w x y z)
-apply (subgoal_tac "u*y + x*y = w*y + v*y  &  u*z + x*z = w*z + v*z")
-apply (simp add: mult_ac)
-apply (simp add: add_mult_distrib [symmetric])
-done
-
-
-lemma mult:
-     "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
-      Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
-by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2
-              UN_equiv_class2 [OF equiv_intrel equiv_intrel])
-
-lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
-by (cases z, cases w, simp add: minus mult add_ac)
-
-lemma zmult_commute: "(z::int) * w = w * z"
-by (cases z, cases w, simp add: mult add_ac mult_ac)
-
-lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
-by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac)
-
-lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
-by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac)
-
-lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
-by (simp add: zmult_commute [of w] zadd_zmult_distrib)
-
-lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
-by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
-
-lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
-by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
-
-lemmas int_distrib =
-  zadd_zmult_distrib zadd_zmult_distrib2
-  zdiff_zmult_distrib zdiff_zmult_distrib2
-
-lemma int_mult: "int (m * n) = (int m) * (int n)"
-by (simp add: int_def mult)
-
-text{*Compatibility binding*}
-lemmas zmult_int = int_mult [symmetric]
-
-lemma zmult_1: "(1::int) * z = z"
-by (cases z, simp add: One_int_def int_def mult)
-
-lemma zmult_1_right: "z * (1::int) = z"
-by (rule trans [OF zmult_commute zmult_1])
-
-
-text{*The integers form a @{text comm_ring_1}*}
-instance int :: comm_ring_1
-proof
-  fix i j k :: int
-  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
-  show "i + j = j + i" by (simp add: zadd_commute)
-  show "0 + i = i" by (rule zadd_0)
-  show "- i + i = 0" by (rule zadd_zminus_inverse2)
-  show "i - j = i + (-j)" by (simp add: diff_int_def)
-  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
-  show "i * j = j * i" by (rule zmult_commute)
-  show "1 * i = i" by (rule zmult_1) 
-  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
-  show "0 \<noteq> (1::int)"
-    by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
-qed
-
-
-subsection{*The @{text "\<le>"} Ordering*}
-
-lemma le:
-  "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
-by (force simp add: le_int_def)
-
-lemma zle_refl: "w \<le> (w::int)"
-by (cases w, simp add: le)
-
-lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)"
-by (cases i, cases j, cases k, simp add: le)
-
-lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)"
-by (cases w, cases z, simp add: le)
-
-instance int :: order
-  by intro_classes
-    (assumption |
-      rule zle_refl zle_trans zle_anti_sym less_int_def [THEN meta_eq_to_obj_eq])+
-
-lemma zle_linear: "(z::int) \<le> w \<or> w \<le> z"
-by (cases z, cases w) (simp add: le linorder_linear)
-
-instance int :: linorder
-  by intro_classes (rule zle_linear)
-
-lemmas zless_linear = linorder_less_linear [where 'a = int]
-
-
-lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)"
-by (simp add: Zero_int_def)
-
-lemma zless_int [simp]: "(int m < int n) = (m<n)"
-by (simp add: le add int_def linorder_not_le [symmetric]) 
-
-lemma int_less_0_conv [simp]: "~ (int k < 0)"
-by (simp add: Zero_int_def)
-
-lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)"
-by (simp add: Zero_int_def)
-
-lemma int_0_less_1: "0 < (1::int)"
-by (simp only: Zero_int_def One_int_def One_nat_def zless_int)
-
-lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)"
-by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq)
-
-lemma zle_int [simp]: "(int m \<le> int n) = (m\<le>n)"
-by (simp add: linorder_not_less [symmetric])
-
-lemma zero_zle_int [simp]: "(0 \<le> int n)"
-by (simp add: Zero_int_def)
-
-lemma int_le_0_conv [simp]: "(int n \<le> 0) = (n = 0)"
-by (simp add: Zero_int_def)
-
-lemma int_0 [simp]: "int 0 = (0::int)"
-by (simp add: Zero_int_def)
-
-lemma int_1 [simp]: "int 1 = 1"
-by (simp add: One_int_def)
-
-lemma int_Suc0_eq_1: "int (Suc 0) = 1"
-by (simp add: One_int_def One_nat_def)
-
-lemma int_Suc: "int (Suc m) = 1 + (int m)"
-by (simp add: One_int_def zadd_int)
-
-
-subsection{*Monotonicity results*}
-
-lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)"
-by (cases i, cases j, cases k, simp add: le add)
-
-lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)"
-apply (cases i, cases j, cases k)
-apply (simp add: linorder_not_le [where 'a = int, symmetric]
-                 linorder_not_le [where 'a = nat]  le add)
-done
-
-lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)"
-by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono])
-
-
-subsection{*Strict Monotonicity of Multiplication*}
-
-text{*strict, in 1st argument; proof is by induction on k>0*}
-lemma zmult_zless_mono2_lemma:
-     "i<j ==> 0<k ==> int k * i < int k * j"
-apply (induct "k", simp)
-apply (simp add: int_Suc)
-apply (case_tac "k=0")
-apply (simp_all add: zadd_zmult_distrib int_Suc0_eq_1 order_le_less)
-apply (simp add: zadd_zless_mono int_Suc0_eq_1 order_le_less)
-done
-
-lemma zero_le_imp_eq_int: "0 \<le> k ==> \<exists>n. k = int n"
-apply (cases k)
-apply (auto simp add: le add int_def Zero_int_def)
-apply (rule_tac x="x-y" in exI, simp)
-done
-
-lemma zmult_zless_mono2: "[| i<j;  (0::int) < k |] ==> k*i < k*j"
-apply (frule order_less_imp_le [THEN zero_le_imp_eq_int])
-apply (auto simp add: zmult_zless_mono2_lemma)
-done
-
-instance int :: minus
-  zabs_def: "\<bar>i\<Colon>int\<bar> \<equiv> if i < 0 then - i else i" ..
-
-instance int :: distrib_lattice
-  "inf \<equiv> min"
-  "sup \<equiv> max"
-  by intro_classes
-    (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1)
-
-text{*The integers form an ordered @{text comm_ring_1}*}
-instance int :: ordered_idom
-proof
-  fix i j k :: int
-  show "i \<le> j ==> k + i \<le> k + j" by (rule zadd_left_mono)
-  show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2)
-  show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def)
-qed
-
-lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
-apply (cases w, cases z) 
-apply (simp add: linorder_not_le [symmetric] le int_def add One_int_def)
-done
-
-subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
-
-definition
-  nat :: "int \<Rightarrow> nat"
-where
-  [code func del]: "nat z = contents (\<Union>(x, y) \<in> Rep_Integ z. {x-y})"
-
-lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
-proof -
-  have "(\<lambda>(x,y). {x-y}) respects intrel"
-    by (simp add: congruent_def) arith
-  thus ?thesis
-    by (simp add: nat_def UN_equiv_class [OF equiv_intrel])
-qed
-
-lemma nat_int [simp]: "nat(int n) = n"
-by (simp add: nat int_def) 
-
-lemma nat_zero [simp]: "nat 0 = 0"
-by (simp only: Zero_int_def nat_int)
-
-lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"
-by (cases z, simp add: nat le int_def Zero_int_def)
-
-corollary nat_0_le: "0 \<le> z ==> int (nat z) = z"
-by simp
-
-lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0"
-by (cases z, simp add: nat le int_def Zero_int_def)
-
-lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
-apply (cases w, cases z) 
-apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith)
-done
-
-text{*An alternative condition is @{term "0 \<le> w"} *}
-corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
-by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
-
-corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)"
-by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) 
-
-lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)"
-apply (cases w, cases z) 
-apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
-done
-
-lemma nonneg_eq_int: "[| 0 \<le> z;  !!m. z = int m ==> P |] ==> P"
-by (blast dest: nat_0_le sym)
-
-lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
-by (cases w, simp add: nat le int_def Zero_int_def, arith)
-
-corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
-by (simp only: eq_commute [of m] nat_eq_iff) 
-
-lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)"
-apply (cases w)
-apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith)
-done
-
-lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)"
-by (auto simp add: nat_eq_iff2)
-
-lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
-by (insert zless_nat_conj [of 0], auto)
-
-lemma nat_add_distrib:
-     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
-by (cases z, cases z', simp add: nat add le int_def Zero_int_def)
-
-lemma nat_diff_distrib:
-     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
-by (cases z, cases z', 
-    simp add: nat add minus diff_minus le int_def Zero_int_def)
-
-
-lemma nat_zminus_int [simp]: "nat (- (int n)) = 0"
-by (simp add: int_def minus nat Zero_int_def) 
-
-lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)"
-by (cases z, simp add: nat le int_def  linorder_not_le [symmetric], arith)
-
-
-subsection{*Lemmas about the Function @{term int} and Orderings*}
-
-lemma negative_zless_0: "- (int (Suc n)) < 0"
-by (simp add: order_less_le)
-
-lemma negative_zless [iff]: "- (int (Suc n)) < int m"
-by (rule negative_zless_0 [THEN order_less_le_trans], simp)
-
-lemma negative_zle_0: "- int n \<le> 0"
-by (simp add: minus_le_iff)
-
-lemma negative_zle [iff]: "- int n \<le> int m"
-by (rule order_trans [OF negative_zle_0 zero_zle_int])
-
-lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))"
-by (subst le_minus_iff, simp)
-
-lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)"
-by (simp add: int_def le minus Zero_int_def) 
-
-lemma not_int_zless_negative [simp]: "~ (int n < - int m)"
-by (simp add: linorder_not_less)
-
-lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)"
-by (force simp add: order_eq_iff [of "- int n"] int_zle_neg)
-
-lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)"
-proof (cases w, cases z, simp add: le add int_def)
-  fix a b c d
-  assume "w = Abs_Integ (intrel `` {(a,b)})" "z = Abs_Integ (intrel `` {(c,d)})"
-  show "(a+d \<le> c+b) = (\<exists>n. c+b = a+n+d)"
-  proof
-    assume "a + d \<le> c + b" 
-    thus "\<exists>n. c + b = a + n + d" 
-      by (auto intro!: exI [where x="c+b - (a+d)"])
-  next    
-    assume "\<exists>n. c + b = a + n + d" 
-    then obtain n where "c + b = a + n + d" ..
-    thus "a + d \<le> c + b" by arith
-  qed
-qed
-
-lemma abs_int_eq [simp]: "abs (int m) = int m"
-by (simp add: abs_if)
-
-text{*This version is proved for all ordered rings, not just integers!
-      It is proved here because attribute @{text arith_split} is not available
-      in theory @{text Ring_and_Field}.
-      But is it really better than just rewriting with @{text abs_if}?*}
-lemma abs_split [arith_split]:
-     "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
-by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
-
-
-subsection {* Constants @{term neg} and @{term iszero} *}
-
-definition
-  neg  :: "'a\<Colon>ordered_idom \<Rightarrow> bool"
-where
-  [code inline]: "neg Z \<longleftrightarrow> Z < 0"
-
-definition (*for simplifying equalities*)
-  iszero :: "'a\<Colon>comm_semiring_1_cancel \<Rightarrow> bool"
-where
-  "iszero z \<longleftrightarrow> z = 0"
-
-lemma not_neg_int [simp]: "~ neg(int n)"
-by (simp add: neg_def)
-
-lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))"
-by (simp add: neg_def neg_less_0_iff_less)
-
-lemmas neg_eq_less_0 = neg_def
-
-lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)"
-by (simp add: neg_def linorder_not_less)
-
-
-subsection{*To simplify inequalities when Numeral1 can get simplified to 1*}
-
-lemma not_neg_0: "~ neg 0"
-by (simp add: One_int_def neg_def)
-
-lemma not_neg_1: "~ neg 1"
-by (simp add: neg_def linorder_not_less zero_le_one)
-
-lemma iszero_0: "iszero 0"
-by (simp add: iszero_def)
-
-lemma not_iszero_1: "~ iszero 1"
-by (simp add: iszero_def eq_commute)
-
-lemma neg_nat: "neg z ==> nat z = 0"
-by (simp add: neg_def order_less_imp_le) 
-
-lemma not_neg_nat: "~ neg z ==> int (nat z) = z"
-by (simp add: linorder_not_less neg_def)
-
-
-subsection{*The Set of Natural Numbers*}
-
-constdefs
-  Nats  :: "'a::semiring_1_cancel set"
-  "Nats == range of_nat"
-
-notation (xsymbols)
-  Nats  ("\<nat>")
-
-lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats"
-by (simp add: Nats_def)
-
-lemma Nats_0 [simp]: "0 \<in> Nats"
-apply (simp add: Nats_def)
-apply (rule range_eqI)
-apply (rule of_nat_0 [symmetric])
-done
-
-lemma Nats_1 [simp]: "1 \<in> Nats"
-apply (simp add: Nats_def)
-apply (rule range_eqI)
-apply (rule of_nat_1 [symmetric])
-done
-
-lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats"
-apply (auto simp add: Nats_def)
-apply (rule range_eqI)
-apply (rule of_nat_add [symmetric])
-done
-
-lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats"
-apply (auto simp add: Nats_def)
-apply (rule range_eqI)
-apply (rule of_nat_mult [symmetric])
-done
-
-text{*Agreement with the specific embedding for the integers*}
-lemma int_eq_of_nat: "int = (of_nat :: nat => int)"
-proof
-  fix n
-  show "int n = of_nat n"  by (induct n, simp_all add: int_Suc add_ac)
-qed
-
-lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)"
-proof
-  fix n
-  show "of_nat n = id n"  by (induct n, simp_all)
-qed
-
-
-subsection{*Embedding of the Integers into any @{text ring_1}:
-@{term of_int}*}
-
-constdefs
-   of_int :: "int => 'a::ring_1"
-   "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
-
-
-lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
-proof -
-  have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
-    by (simp add: congruent_def compare_rls of_nat_add [symmetric]
-            del: of_nat_add) 
-  thus ?thesis
-    by (simp add: of_int_def UN_equiv_class [OF equiv_intrel])
-qed
-
-lemma of_int_0 [simp]: "of_int 0 = 0"
-by (simp add: of_int Zero_int_def int_def)
-
-lemma of_int_1 [simp]: "of_int 1 = 1"
-by (simp add: of_int One_int_def int_def)
-
-lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z"
-by (cases w, cases z, simp add: compare_rls of_int add)
-
-lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)"
-by (cases z, simp add: compare_rls of_int minus)
-
-lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z"
-by (simp add: diff_minus)
-
-lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
-apply (cases w, cases z)
-apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib
-                 mult add_ac)
-done
-
-lemma of_int_le_iff [simp]:
-     "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)"
-apply (cases w)
-apply (cases z)
-apply (simp add: compare_rls of_int le diff_int_def add minus
-                 of_nat_add [symmetric]   del: of_nat_add)
-done
-
-text{*Special cases where either operand is zero*}
-lemmas of_int_0_le_iff [simp] = of_int_le_iff [of 0, simplified]
-lemmas of_int_le_0_iff [simp] = of_int_le_iff [of _ 0, simplified]
-
-
-lemma of_int_less_iff [simp]:
-     "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)"
-by (simp add: linorder_not_le [symmetric])
-
-text{*Special cases where either operand is zero*}
-lemmas of_int_0_less_iff [simp] = of_int_less_iff [of 0, simplified]
-lemmas of_int_less_0_iff [simp] = of_int_less_iff [of _ 0, simplified]
-
-text{*Class for unital rings with characteristic zero.
- Includes non-ordered rings like the complex numbers.*}
-axclass ring_char_0 < ring_1
-  of_int_inject: "of_int w = of_int z ==> w = z"
-
-lemma of_int_eq_iff [simp]:
-     "(of_int w = (of_int z::'a::ring_char_0)) = (w = z)"
-by (safe elim!: of_int_inject)
-
-text{*Every @{text ordered_idom} has characteristic zero.*}
-instance ordered_idom < ring_char_0
-by intro_classes (simp add: order_eq_iff)
-
-text{*Special cases where either operand is zero*}
-lemmas of_int_0_eq_iff [simp] = of_int_eq_iff [of 0, simplified]
-lemmas of_int_eq_0_iff [simp] = of_int_eq_iff [of _ 0, simplified]
-
-lemma of_int_eq_id [simp]: "of_int = (id :: int => int)"
-proof
-  fix z
-  show "of_int z = id z"  
-    by (cases z)
-      (simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus)
-qed
-
-
-subsection{*The Set of Integers*}
-
-constdefs
-  Ints  :: "'a::ring_1 set"
-  "Ints == range of_int"
-
-notation (xsymbols)
-  Ints  ("\<int>")
-
-lemma Ints_0 [simp]: "0 \<in> Ints"
-apply (simp add: Ints_def)
-apply (rule range_eqI)
-apply (rule of_int_0 [symmetric])
-done
-
-lemma Ints_1 [simp]: "1 \<in> Ints"
-apply (simp add: Ints_def)
-apply (rule range_eqI)
-apply (rule of_int_1 [symmetric])
-done
-
-lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints"
-apply (auto simp add: Ints_def)
-apply (rule range_eqI)
-apply (rule of_int_add [symmetric])
-done
-
-lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints"
-apply (auto simp add: Ints_def)
-apply (rule range_eqI)
-apply (rule of_int_minus [symmetric])
-done
-
-lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints"
-apply (auto simp add: Ints_def)
-apply (rule range_eqI)
-apply (rule of_int_diff [symmetric])
-done
-
-lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints"
-apply (auto simp add: Ints_def)
-apply (rule range_eqI)
-apply (rule of_int_mult [symmetric])
-done
-
-text{*Collapse nested embeddings*}
-lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n"
-by (induct n, auto)
-
-lemma of_int_int_eq [simp]: "of_int (int n) = of_nat n"
-by (simp add: int_eq_of_nat)
-
-lemma Ints_cases [cases set: Ints]:
-  assumes "q \<in> \<int>"
-  obtains (of_int) z where "q = of_int z"
-  unfolding Ints_def
-proof -
-  from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def .
-  then obtain z where "q = of_int z" ..
-  then show thesis ..
-qed
-
-lemma Ints_induct [case_names of_int, induct set: Ints]:
-  "q \<in> \<int> ==> (!!z. P (of_int z)) ==> P q"
-  by (rule Ints_cases) auto
-
-
-(* int (Suc n) = 1 + int n *)
-
-
-
-subsection{*More Properties of @{term setsum} and  @{term setprod}*}
-
-text{*By Jeremy Avigad*}
-
-
-lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
-  apply (cases "finite A")
-  apply (erule finite_induct, auto)
-  done
-
-lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))"
-  apply (cases "finite A")
-  apply (erule finite_induct, auto)
-  done
-
-lemma int_setsum: "int (setsum f A) = (\<Sum>x\<in>A. int(f x))"
-  by (simp add: int_eq_of_nat of_nat_setsum)
-
-lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))"
-  apply (cases "finite A")
-  apply (erule finite_induct, auto)
-  done
-
-lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
-  apply (cases "finite A")
-  apply (erule finite_induct, auto)
-  done
-
-lemma int_setprod: "int (setprod f A) = (\<Prod>x\<in>A. int(f x))"
-  by (simp add: int_eq_of_nat of_nat_setprod)
-
-lemma setprod_nonzero_nat:
-    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0"
-  by (rule setprod_nonzero, auto)
-
-lemma setprod_zero_eq_nat:
-    "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)"
-  by (rule setprod_zero_eq, auto)
-
-lemma setprod_nonzero_int:
-    "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0"
-  by (rule setprod_nonzero, auto)
-
-lemma setprod_zero_eq_int:
-    "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)"
-  by (rule setprod_zero_eq, auto)
-
-
-subsection {* Further properties *}
-
-text{*Now we replace the case analysis rule by a more conventional one:
-whether an integer is negative or not.*}
-
-lemma zless_iff_Suc_zadd:
-    "(w < z) = (\<exists>n. z = w + int(Suc n))"
-apply (cases z, cases w)
-apply (auto simp add: le add int_def linorder_not_le [symmetric]) 
-apply (rename_tac a b c d) 
-apply (rule_tac x="a+d - Suc(c+b)" in exI) 
-apply arith
-done
-
-lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))"
-apply (cases x)
-apply (auto simp add: le minus Zero_int_def int_def order_less_le) 
-apply (rule_tac x="y - Suc x" in exI, arith)
-done
-
-theorem int_cases [cases type: int, case_names nonneg neg]:
-     "[|!! n. z = int n ==> P;  !! n. z =  - (int (Suc n)) ==> P |] ==> P"
-apply (cases "z < 0", blast dest!: negD)
-apply (simp add: linorder_not_less)
-apply (blast dest: nat_0_le [THEN sym])
-done
-
-theorem int_induct [induct type: int, case_names nonneg neg]:
-     "[|!! n. P (int n);  !!n. P (- (int (Suc n))) |] ==> P z"
-  by (cases z) auto
-
-text{*Contributed by Brian Huffman*}
-theorem int_diff_cases [case_names diff]:
-assumes prem: "!!m n. z = int m - int n ==> P" shows "P"
- apply (rule_tac z=z in int_cases)
-  apply (rule_tac m=n and n=0 in prem, simp)
- apply (rule_tac m=0 and n="Suc n" in prem, simp)
-done
-
-lemma of_nat_nat: "0 \<le> z ==> of_nat (nat z) = of_int z"
-apply (cases z)
-apply (simp_all add: not_zle_0_negative del: int_Suc)
-done
-
-lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq]
-
-lemmas [simp] = int_Suc
-
-
-subsection {* Legacy ML bindings *}
-
-ML {*
-val of_nat_0 = @{thm of_nat_0};
-val of_nat_1 = @{thm of_nat_1};
-val of_nat_Suc = @{thm of_nat_Suc};
-val of_nat_add = @{thm of_nat_add};
-val of_nat_mult = @{thm of_nat_mult};
-val of_int_0 = @{thm of_int_0};
-val of_int_1 = @{thm of_int_1};
-val of_int_add = @{thm of_int_add};
-val of_int_mult = @{thm of_int_mult};
-val int_eq_of_nat = @{thm int_eq_of_nat};
-val zle_int = @{thm zle_int};
-val int_int_eq = @{thm int_int_eq};
-val diff_int_def = @{thm diff_int_def};
-val zadd_ac = @{thms zadd_ac};
-val zless_int = @{thm zless_int};
-val zadd_int = @{thm zadd_int};
-val zmult_int = @{thm zmult_int};
-val nat_0_le = @{thm nat_0_le};
-val int_0 = @{thm int_0};
-val int_1 = @{thm int_1};
-val abs_split = @{thm abs_split};
-*}
-
-end
--- a/src/HOL/Integ/IntDiv.thy	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1406 +0,0 @@
-(*  Title:      HOL/IntDiv.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
-
-*)
-
-header{*The Division Operators div and mod; the Divides Relation dvd*}
-
-theory IntDiv
-imports IntArith "../Divides" "../FunDef"
-begin
-
-declare zless_nat_conj [simp]
-
-constdefs
-  quorem :: "(int*int) * (int*int) => bool"
-    --{*definition of quotient and remainder*}
-    [code func]: "quorem == %((a,b), (q,r)).
-                      a = b*q + r &
-                      (if 0 < b then 0\<le>r & r<b else b<r & r \<le> 0)"
-
-  adjust :: "[int, int*int] => int*int"
-    --{*for the division algorithm*}
-    [code func]: "adjust b == %(q,r). if 0 \<le> r-b then (2*q + 1, r-b)
-                         else (2*q, r)"
-
-text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
-function
-  posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
-where
-  "posDivAlg a b =
-     (if (a<b | b\<le>0) then (0,a)
-        else adjust b (posDivAlg a (2*b)))"
-by auto
-termination by (relation "measure (%(a,b). nat(a - b + 1))") auto
-
-text{*algorithm for the case @{text "a<0, b>0"}*}
-function
-  negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
-where
-  "negDivAlg a b  =
-     (if (0\<le>a+b | b\<le>0) then (-1,a+b)
-      else adjust b (negDivAlg a (2*b)))"
-by auto
-termination by (relation "measure (%(a,b). nat(- a - b))") auto
-
-text{*algorithm for the general case @{term "b\<noteq>0"}*}
-constdefs
-  negateSnd :: "int*int => int*int"
-    [code func]: "negateSnd == %(q,r). (q,-r)"
-
-definition
-  divAlg :: "int \<times> int \<Rightarrow> int \<times> int"
-    --{*The full division algorithm considers all possible signs for a, b
-       including the special case @{text "a=0, b<0"} because 
-       @{term negDivAlg} requires @{term "a<0"}.*}
-where
-  "divAlg = (\<lambda>(a, b). (if 0\<le>a then
-                  if 0\<le>b then posDivAlg a b
-                  else if a=0 then (0, 0)
-                       else negateSnd (negDivAlg (-a) (-b))
-               else 
-                  if 0<b then negDivAlg a b
-                  else negateSnd (posDivAlg (-a) (-b))))"
-
-instance int :: Divides.div
-  div_def: "a div b == fst (divAlg (a, b))"
-  mod_def: "a mod b == snd (divAlg (a, b))" ..
-
-lemma divAlg_mod_div:
-  "divAlg (p, q) = (p div q, p mod q)"
-  by (auto simp add: div_def mod_def)
-
-text{*
-Here is the division algorithm in ML:
-
-\begin{verbatim}
-    fun posDivAlg (a,b) =
-      if a<b then (0,a)
-      else let val (q,r) = posDivAlg(a, 2*b)
-	       in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
-	   end
-
-    fun negDivAlg (a,b) =
-      if 0\<le>a+b then (~1,a+b)
-      else let val (q,r) = negDivAlg(a, 2*b)
-	       in  if 0\<le>r-b then (2*q+1, r-b) else (2*q, r)
-	   end;
-
-    fun negateSnd (q,r:int) = (q,~r);
-
-    fun divAlg (a,b) = if 0\<le>a then 
-			  if b>0 then posDivAlg (a,b) 
-			   else if a=0 then (0,0)
-				else negateSnd (negDivAlg (~a,~b))
-		       else 
-			  if 0<b then negDivAlg (a,b)
-			  else        negateSnd (posDivAlg (~a,~b));
-\end{verbatim}
-*}
-
-
-
-subsection{*Uniqueness and Monotonicity of Quotients and Remainders*}
-
-lemma unique_quotient_lemma:
-     "[| b*q' + r'  \<le> b*q + r;  0 \<le> r';  r' < b;  r < b |]  
-      ==> q' \<le> (q::int)"
-apply (subgoal_tac "r' + b * (q'-q) \<le> r")
- prefer 2 apply (simp add: right_diff_distrib)
-apply (subgoal_tac "0 < b * (1 + q - q') ")
-apply (erule_tac [2] order_le_less_trans)
- prefer 2 apply (simp add: right_diff_distrib right_distrib)
-apply (subgoal_tac "b * q' < b * (1 + q) ")
- prefer 2 apply (simp add: right_diff_distrib right_distrib)
-apply (simp add: mult_less_cancel_left)
-done
-
-lemma unique_quotient_lemma_neg:
-     "[| b*q' + r' \<le> b*q + r;  r \<le> 0;  b < r;  b < r' |]  
-      ==> q \<le> (q'::int)"
-by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, 
-    auto)
-
-lemma unique_quotient:
-     "[| quorem ((a,b), (q,r));  quorem ((a,b), (q',r'));  b \<noteq> 0 |]  
-      ==> q = q'"
-apply (simp add: quorem_def linorder_neq_iff split: split_if_asm)
-apply (blast intro: order_antisym
-             dest: order_eq_refl [THEN unique_quotient_lemma] 
-             order_eq_refl [THEN unique_quotient_lemma_neg] sym)+
-done
-
-
-lemma unique_remainder:
-     "[| quorem ((a,b), (q,r));  quorem ((a,b), (q',r'));  b \<noteq> 0 |]  
-      ==> r = r'"
-apply (subgoal_tac "q = q'")
- apply (simp add: quorem_def)
-apply (blast intro: unique_quotient)
-done
-
-
-subsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*}
-
-text{*And positive divisors*}
-
-lemma adjust_eq [simp]:
-     "adjust b (q,r) = 
-      (let diff = r-b in  
-	if 0 \<le> diff then (2*q + 1, diff)   
-                     else (2*q, r))"
-by (simp add: Let_def adjust_def)
-
-declare posDivAlg.simps [simp del]
-
-text{*use with a simproc to avoid repeatedly proving the premise*}
-lemma posDivAlg_eqn:
-     "0 < b ==>  
-      posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))"
-by (rule posDivAlg.simps [THEN trans], simp)
-
-text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
-theorem posDivAlg_correct:
-  assumes "0 \<le> a" and "0 < b"
-  shows "quorem ((a, b), posDivAlg a b)"
-using prems apply (induct a b rule: posDivAlg.induct)
-apply auto
-apply (simp add: quorem_def)
-apply (subst posDivAlg_eqn, simp add: right_distrib)
-apply (case_tac "a < b")
-apply simp_all
-apply (erule splitE)
-apply (auto simp add: right_distrib Let_def)
-done
-
-
-subsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*}
-
-text{*And positive divisors*}
-
-declare negDivAlg.simps [simp del]
-
-text{*use with a simproc to avoid repeatedly proving the premise*}
-lemma negDivAlg_eqn:
-     "0 < b ==>  
-      negDivAlg a b =       
-       (if 0\<le>a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))"
-by (rule negDivAlg.simps [THEN trans], simp)
-
-(*Correctness of negDivAlg: it computes quotients correctly
-  It doesn't work if a=0 because the 0/b equals 0, not -1*)
-lemma negDivAlg_correct:
-  assumes "a < 0" and "b > 0"
-  shows "quorem ((a, b), negDivAlg a b)"
-using prems apply (induct a b rule: negDivAlg.induct)
-apply (auto simp add: linorder_not_le)
-apply (simp add: quorem_def)
-apply (subst negDivAlg_eqn, assumption)
-apply (case_tac "a + b < (0\<Colon>int)")
-apply simp_all
-apply (erule splitE)
-apply (auto simp add: right_distrib Let_def)
-done
-
-
-subsection{*Existence Shown by Proving the Division Algorithm to be Correct*}
-
-(*the case a=0*)
-lemma quorem_0: "b \<noteq> 0 ==> quorem ((0,b), (0,0))"
-by (auto simp add: quorem_def linorder_neq_iff)
-
-lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)"
-by (subst posDivAlg.simps, auto)
-
-lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)"
-by (subst negDivAlg.simps, auto)
-
-lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)"
-by (simp add: negateSnd_def)
-
-lemma quorem_neg: "quorem ((-a,-b), qr) ==> quorem ((a,b), negateSnd qr)"
-by (auto simp add: split_ifs quorem_def)
-
-lemma divAlg_correct: "b \<noteq> 0 ==> quorem ((a,b), divAlg (a, b))"
-by (force simp add: linorder_neq_iff quorem_0 divAlg_def quorem_neg
-                    posDivAlg_correct negDivAlg_correct)
-
-text{*Arbitrary definitions for division by zero.  Useful to simplify 
-    certain equations.*}
-
-lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a"
-by (simp add: div_def mod_def divAlg_def posDivAlg.simps)  
-
-
-text{*Basic laws about division and remainder*}
-
-lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)"
-apply (case_tac "b = 0", simp)
-apply (cut_tac a = a and b = b in divAlg_correct)
-apply (auto simp add: quorem_def div_def mod_def)
-done
-
-lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k"
-by(simp add: zmod_zdiv_equality[symmetric])
-
-lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k"
-by(simp add: mult_commute zmod_zdiv_equality[symmetric])
-
-text {* Tool setup *}
-
-ML_setup {*
-local 
-
-structure CancelDivMod = CancelDivModFun(
-struct
-  val div_name = @{const_name Divides.div};
-  val mod_name = @{const_name Divides.mod};
-  val mk_binop = HOLogic.mk_binop;
-  val mk_sum = Int_Numeral_Simprocs.mk_sum HOLogic.intT;
-  val dest_sum = Int_Numeral_Simprocs.dest_sum;
-  val div_mod_eqs =
-    map mk_meta_eq [@{thm zdiv_zmod_equality},
-      @{thm zdiv_zmod_equality2}];
-  val trans = trans;
-  val prove_eq_sums =
-    let
-      val simps = diff_int_def :: Int_Numeral_Simprocs.add_0s @ zadd_ac
-    in NatArithUtils.prove_conv all_tac (NatArithUtils.simp_all_tac simps) end;
-end)
-
-in
-
-val cancel_zdiv_zmod_proc = NatArithUtils.prep_simproc
-  ("cancel_zdiv_zmod", ["(m::int) + n"], K CancelDivMod.proc)
-
-end;
-
-Addsimprocs [cancel_zdiv_zmod_proc]
-*}
-
-lemma pos_mod_conj : "(0::int) < b ==> 0 \<le> a mod b & a mod b < b"
-apply (cut_tac a = a and b = b in divAlg_correct)
-apply (auto simp add: quorem_def mod_def)
-done
-
-lemmas pos_mod_sign  [simp] = pos_mod_conj [THEN conjunct1, standard]
-   and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard]
-
-lemma neg_mod_conj : "b < (0::int) ==> a mod b \<le> 0 & b < a mod b"
-apply (cut_tac a = a and b = b in divAlg_correct)
-apply (auto simp add: quorem_def div_def mod_def)
-done
-
-lemmas neg_mod_sign  [simp] = neg_mod_conj [THEN conjunct1, standard]
-   and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard]
-
-
-
-subsection{*General Properties of div and mod*}
-
-lemma quorem_div_mod: "b \<noteq> 0 ==> quorem ((a, b), (a div b, a mod b))"
-apply (cut_tac a = a and b = b in zmod_zdiv_equality)
-apply (force simp add: quorem_def linorder_neq_iff)
-done
-
-lemma quorem_div: "[| quorem((a,b),(q,r));  b \<noteq> 0 |] ==> a div b = q"
-by (simp add: quorem_div_mod [THEN unique_quotient])
-
-lemma quorem_mod: "[| quorem((a,b),(q,r));  b \<noteq> 0 |] ==> a mod b = r"
-by (simp add: quorem_div_mod [THEN unique_remainder])
-
-lemma div_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a div b = 0"
-apply (rule quorem_div)
-apply (auto simp add: quorem_def)
-done
-
-lemma div_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a div b = 0"
-apply (rule quorem_div)
-apply (auto simp add: quorem_def)
-done
-
-lemma div_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a div b = -1"
-apply (rule quorem_div)
-apply (auto simp add: quorem_def)
-done
-
-(*There is no div_neg_pos_trivial because  0 div b = 0 would supersede it*)
-
-lemma mod_pos_pos_trivial: "[| (0::int) \<le> a;  a < b |] ==> a mod b = a"
-apply (rule_tac q = 0 in quorem_mod)
-apply (auto simp add: quorem_def)
-done
-
-lemma mod_neg_neg_trivial: "[| a \<le> (0::int);  b < a |] ==> a mod b = a"
-apply (rule_tac q = 0 in quorem_mod)
-apply (auto simp add: quorem_def)
-done
-
-lemma mod_pos_neg_trivial: "[| (0::int) < a;  a+b \<le> 0 |] ==> a mod b = a+b"
-apply (rule_tac q = "-1" in quorem_mod)
-apply (auto simp add: quorem_def)
-done
-
-text{*There is no @{text mod_neg_pos_trivial}.*}
-
-
-(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*)
-lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)"
-apply (case_tac "b = 0", simp)
-apply (simp add: quorem_div_mod [THEN quorem_neg, simplified, 
-                                 THEN quorem_div, THEN sym])
-
-done
-
-(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*)
-lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))"
-apply (case_tac "b = 0", simp)
-apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod],
-       auto)
-done
-
-
-subsection{*Laws for div and mod with Unary Minus*}
-
-lemma zminus1_lemma:
-     "quorem((a,b),(q,r))  
-      ==> quorem ((-a,b), (if r=0 then -q else -q - 1),  
-                          (if r=0 then 0 else b-r))"
-by (force simp add: split_ifs quorem_def linorder_neq_iff right_diff_distrib)
-
-
-lemma zdiv_zminus1_eq_if:
-     "b \<noteq> (0::int)  
-      ==> (-a) div b =  
-          (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
-by (blast intro: quorem_div_mod [THEN zminus1_lemma, THEN quorem_div])
-
-lemma zmod_zminus1_eq_if:
-     "(-a::int) mod b = (if a mod b = 0 then 0 else  b - (a mod b))"
-apply (case_tac "b = 0", simp)
-apply (blast intro: quorem_div_mod [THEN zminus1_lemma, THEN quorem_mod])
-done
-
-lemma zdiv_zminus2: "a div (-b) = (-a::int) div b"
-by (cut_tac a = "-a" in zdiv_zminus_zminus, auto)
-
-lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)"
-by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto)
-
-lemma zdiv_zminus2_eq_if:
-     "b \<noteq> (0::int)  
-      ==> a div (-b) =  
-          (if a mod b = 0 then - (a div b) else  - (a div b) - 1)"
-by (simp add: zdiv_zminus1_eq_if zdiv_zminus2)
-
-lemma zmod_zminus2_eq_if:
-     "a mod (-b::int) = (if a mod b = 0 then 0 else  (a mod b) - b)"
-by (simp add: zmod_zminus1_eq_if zmod_zminus2)
-
-
-subsection{*Division of a Number by Itself*}
-
-lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \<le> q"
-apply (subgoal_tac "0 < a*q")
- apply (simp add: zero_less_mult_iff, arith)
-done
-
-lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \<le> r |] ==> q \<le> 1"
-apply (subgoal_tac "0 \<le> a* (1-q) ")
- apply (simp add: zero_le_mult_iff)
-apply (simp add: right_diff_distrib)
-done
-
-lemma self_quotient: "[| quorem((a,a),(q,r));  a \<noteq> (0::int) |] ==> q = 1"
-apply (simp add: split_ifs quorem_def linorder_neq_iff)
-apply (rule order_antisym, safe, simp_all)
-apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1)
-apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2)
-apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+
-done
-
-lemma self_remainder: "[| quorem((a,a),(q,r));  a \<noteq> (0::int) |] ==> r = 0"
-apply (frule self_quotient, assumption)
-apply (simp add: quorem_def)
-done
-
-lemma zdiv_self [simp]: "a \<noteq> 0 ==> a div a = (1::int)"
-by (simp add: quorem_div_mod [THEN self_quotient])
-
-(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *)
-lemma zmod_self [simp]: "a mod a = (0::int)"
-apply (case_tac "a = 0", simp)
-apply (simp add: quorem_div_mod [THEN self_remainder])
-done
-
-
-subsection{*Computation of Division and Remainder*}
-
-lemma zdiv_zero [simp]: "(0::int) div b = 0"
-by (simp add: div_def divAlg_def)
-
-lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1"
-by (simp add: div_def divAlg_def)
-
-lemma zmod_zero [simp]: "(0::int) mod b = 0"
-by (simp add: mod_def divAlg_def)
-
-lemma zdiv_minus1: "(0::int) < b ==> -1 div b = -1"
-by (simp add: div_def divAlg_def)
-
-lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1"
-by (simp add: mod_def divAlg_def)
-
-text{*a positive, b positive *}
-
-lemma div_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a div b = fst (posDivAlg a b)"
-by (simp add: div_def divAlg_def)
-
-lemma mod_pos_pos: "[| 0 < a;  0 \<le> b |] ==> a mod b = snd (posDivAlg a b)"
-by (simp add: mod_def divAlg_def)
-
-text{*a negative, b positive *}
-
-lemma div_neg_pos: "[| a < 0;  0 < b |] ==> a div b = fst (negDivAlg a b)"
-by (simp add: div_def divAlg_def)
-
-lemma mod_neg_pos: "[| a < 0;  0 < b |] ==> a mod b = snd (negDivAlg a b)"
-by (simp add: mod_def divAlg_def)
-
-text{*a positive, b negative *}
-
-lemma div_pos_neg:
-     "[| 0 < a;  b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))"
-by (simp add: div_def divAlg_def)
-
-lemma mod_pos_neg:
-     "[| 0 < a;  b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))"
-by (simp add: mod_def divAlg_def)
-
-text{*a negative, b negative *}
-
-lemma div_neg_neg:
-     "[| a < 0;  b \<le> 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))"
-by (simp add: div_def divAlg_def)
-
-lemma mod_neg_neg:
-     "[| a < 0;  b \<le> 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))"
-by (simp add: mod_def divAlg_def)
-
-text {*Simplify expresions in which div and mod combine numerical constants*}
-
-lemmas div_pos_pos_number_of [simp] =
-    div_pos_pos [of "number_of v" "number_of w", standard]
-
-lemmas div_neg_pos_number_of [simp] =
-    div_neg_pos [of "number_of v" "number_of w", standard]
-
-lemmas div_pos_neg_number_of [simp] =
-    div_pos_neg [of "number_of v" "number_of w", standard]
-
-lemmas div_neg_neg_number_of [simp] =
-    div_neg_neg [of "number_of v" "number_of w", standard]
-
-
-lemmas mod_pos_pos_number_of [simp] =
-    mod_pos_pos [of "number_of v" "number_of w", standard]
-
-lemmas mod_neg_pos_number_of [simp] =
-    mod_neg_pos [of "number_of v" "number_of w", standard]
-
-lemmas mod_pos_neg_number_of [simp] =
-    mod_pos_neg [of "number_of v" "number_of w", standard]
-
-lemmas mod_neg_neg_number_of [simp] =
-    mod_neg_neg [of "number_of v" "number_of w", standard]
-
-
-lemmas posDivAlg_eqn_number_of [simp] =
-    posDivAlg_eqn [of "number_of v" "number_of w", standard]
-
-lemmas negDivAlg_eqn_number_of [simp] =
-    negDivAlg_eqn [of "number_of v" "number_of w", standard]
-
-
-text{*Special-case simplification *}
-
-lemma zmod_1 [simp]: "a mod (1::int) = 0"
-apply (cut_tac a = a and b = 1 in pos_mod_sign)
-apply (cut_tac [2] a = a and b = 1 in pos_mod_bound)
-apply (auto simp del:pos_mod_bound pos_mod_sign)
-done
-
-lemma zdiv_1 [simp]: "a div (1::int) = a"
-by (cut_tac a = a and b = 1 in zmod_zdiv_equality, auto)
-
-lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0"
-apply (cut_tac a = a and b = "-1" in neg_mod_sign)
-apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound)
-apply (auto simp del: neg_mod_sign neg_mod_bound)
-done
-
-lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a"
-by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto)
-
-(** The last remaining special cases for constant arithmetic:
-    1 div z and 1 mod z **)
-
-lemmas div_pos_pos_1_number_of [simp] =
-    div_pos_pos [OF int_0_less_1, of "number_of w", standard]
-
-lemmas div_pos_neg_1_number_of [simp] =
-    div_pos_neg [OF int_0_less_1, of "number_of w", standard]
-
-lemmas mod_pos_pos_1_number_of [simp] =
-    mod_pos_pos [OF int_0_less_1, of "number_of w", standard]
-
-lemmas mod_pos_neg_1_number_of [simp] =
-    mod_pos_neg [OF int_0_less_1, of "number_of w", standard]
-
-
-lemmas posDivAlg_eqn_1_number_of [simp] =
-    posDivAlg_eqn [of concl: 1 "number_of w", standard]
-
-lemmas negDivAlg_eqn_1_number_of [simp] =
-    negDivAlg_eqn [of concl: 1 "number_of w", standard]
-
-
-
-subsection{*Monotonicity in the First Argument (Dividend)*}
-
-lemma zdiv_mono1: "[| a \<le> a';  0 < (b::int) |] ==> a div b \<le> a' div b"
-apply (cut_tac a = a and b = b in zmod_zdiv_equality)
-apply (cut_tac a = a' and b = b in zmod_zdiv_equality)
-apply (rule unique_quotient_lemma)
-apply (erule subst)
-apply (erule subst, simp_all)
-done
-
-lemma zdiv_mono1_neg: "[| a \<le> a';  (b::int) < 0 |] ==> a' div b \<le> a div b"
-apply (cut_tac a = a and b = b in zmod_zdiv_equality)
-apply (cut_tac a = a' and b = b in zmod_zdiv_equality)
-apply (rule unique_quotient_lemma_neg)
-apply (erule subst)
-apply (erule subst, simp_all)
-done
-
-
-subsection{*Monotonicity in the Second Argument (Divisor)*}
-
-lemma q_pos_lemma:
-     "[| 0 \<le> b'*q' + r'; r' < b';  0 < b' |] ==> 0 \<le> (q'::int)"
-apply (subgoal_tac "0 < b'* (q' + 1) ")
- apply (simp add: zero_less_mult_iff)
-apply (simp add: right_distrib)
-done
-
-lemma zdiv_mono2_lemma:
-     "[| b*q + r = b'*q' + r';  0 \<le> b'*q' + r';   
-         r' < b';  0 \<le> r;  0 < b';  b' \<le> b |]   
-      ==> q \<le> (q'::int)"
-apply (frule q_pos_lemma, assumption+) 
-apply (subgoal_tac "b*q < b* (q' + 1) ")
- apply (simp add: mult_less_cancel_left)
-apply (subgoal_tac "b*q = r' - r + b'*q'")
- prefer 2 apply simp
-apply (simp (no_asm_simp) add: right_distrib)
-apply (subst add_commute, rule zadd_zless_mono, arith)
-apply (rule mult_right_mono, auto)
-done
-
-lemma zdiv_mono2:
-     "[| (0::int) \<le> a;  0 < b';  b' \<le> b |] ==> a div b \<le> a div b'"
-apply (subgoal_tac "b \<noteq> 0")
- prefer 2 apply arith
-apply (cut_tac a = a and b = b in zmod_zdiv_equality)
-apply (cut_tac a = a and b = b' in zmod_zdiv_equality)
-apply (rule zdiv_mono2_lemma)
-apply (erule subst)
-apply (erule subst, simp_all)
-done
-
-lemma q_neg_lemma:
-     "[| b'*q' + r' < 0;  0 \<le> r';  0 < b' |] ==> q' \<le> (0::int)"
-apply (subgoal_tac "b'*q' < 0")
- apply (simp add: mult_less_0_iff, arith)
-done
-
-lemma zdiv_mono2_neg_lemma:
-     "[| b*q + r = b'*q' + r';  b'*q' + r' < 0;   
-         r < b;  0 \<le> r';  0 < b';  b' \<le> b |]   
-      ==> q' \<le> (q::int)"
-apply (frule q_neg_lemma, assumption+) 
-apply (subgoal_tac "b*q' < b* (q + 1) ")
- apply (simp add: mult_less_cancel_left)
-apply (simp add: right_distrib)
-apply (subgoal_tac "b*q' \<le> b'*q'")
- prefer 2 apply (simp add: mult_right_mono_neg, arith)
-done
-
-lemma zdiv_mono2_neg:
-     "[| a < (0::int);  0 < b';  b' \<le> b |] ==> a div b' \<le> a div b"
-apply (cut_tac a = a and b = b in zmod_zdiv_equality)
-apply (cut_tac a = a and b = b' in zmod_zdiv_equality)
-apply (rule zdiv_mono2_neg_lemma)
-apply (erule subst)
-apply (erule subst, simp_all)
-done
-
-subsection{*More Algebraic Laws for div and mod*}
-
-text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
-
-lemma zmult1_lemma:
-     "[| quorem((b,c),(q,r));  c \<noteq> 0 |]  
-      ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))"
-by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib)
-
-lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"
-apply (case_tac "c = 0", simp)
-apply (blast intro: quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
-done
-
-lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)"
-apply (case_tac "c = 0", simp)
-apply (blast intro: quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
-done
-
-lemma zmod_zmult1_eq': "(a*b) mod (c::int) = ((a mod c) * b) mod c"
-apply (rule trans)
-apply (rule_tac s = "b*a mod c" in trans)
-apply (rule_tac [2] zmod_zmult1_eq)
-apply (simp_all add: mult_commute)
-done
-
-lemma zmod_zmult_distrib: "(a*b) mod (c::int) = ((a mod c) * (b mod c)) mod c"
-apply (rule zmod_zmult1_eq' [THEN trans])
-apply (rule zmod_zmult1_eq)
-done
-
-lemma zdiv_zmult_self1 [simp]: "b \<noteq> (0::int) ==> (a*b) div b = a"
-by (simp add: zdiv_zmult1_eq)
-
-lemma zdiv_zmult_self2 [simp]: "b \<noteq> (0::int) ==> (b*a) div b = a"
-by (subst mult_commute, erule zdiv_zmult_self1)
-
-lemma zmod_zmult_self1 [simp]: "(a*b) mod b = (0::int)"
-by (simp add: zmod_zmult1_eq)
-
-lemma zmod_zmult_self2 [simp]: "(b*a) mod b = (0::int)"
-by (simp add: mult_commute zmod_zmult1_eq)
-
-lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
-proof
-  assume "m mod d = 0"
-  with zmod_zdiv_equality[of m d] show "EX q::int. m = d*q" by auto
-next
-  assume "EX q::int. m = d*q"
-  thus "m mod d = 0" by auto
-qed
-
-lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1]
-
-text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
-
-lemma zadd1_lemma:
-     "[| quorem((a,c),(aq,ar));  quorem((b,c),(bq,br));  c \<noteq> 0 |]  
-      ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))"
-by (force simp add: split_ifs quorem_def linorder_neq_iff right_distrib)
-
-(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
-lemma zdiv_zadd1_eq:
-     "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
-apply (case_tac "c = 0", simp)
-apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_div)
-done
-
-lemma zmod_zadd1_eq: "(a+b) mod (c::int) = (a mod c + b mod c) mod c"
-apply (case_tac "c = 0", simp)
-apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod] quorem_mod)
-done
-
-lemma mod_div_trivial [simp]: "(a mod b) div b = (0::int)"
-apply (case_tac "b = 0", simp)
-apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial)
-done
-
-lemma mod_mod_trivial [simp]: "(a mod b) mod b = a mod (b::int)"
-apply (case_tac "b = 0", simp)
-apply (force simp add: linorder_neq_iff mod_pos_pos_trivial mod_neg_neg_trivial)
-done
-
-lemma zmod_zadd_left_eq: "(a+b) mod (c::int) = ((a mod c) + b) mod c"
-apply (rule trans [symmetric])
-apply (rule zmod_zadd1_eq, simp)
-apply (rule zmod_zadd1_eq [symmetric])
-done
-
-lemma zmod_zadd_right_eq: "(a+b) mod (c::int) = (a + (b mod c)) mod c"
-apply (rule trans [symmetric])
-apply (rule zmod_zadd1_eq, simp)
-apply (rule zmod_zadd1_eq [symmetric])
-done
-
-lemma zdiv_zadd_self1[simp]: "a \<noteq> (0::int) ==> (a+b) div a = b div a + 1"
-by (simp add: zdiv_zadd1_eq)
-
-lemma zdiv_zadd_self2[simp]: "a \<noteq> (0::int) ==> (b+a) div a = b div a + 1"
-by (simp add: zdiv_zadd1_eq)
-
-lemma zmod_zadd_self1[simp]: "(a+b) mod a = b mod (a::int)"
-apply (case_tac "a = 0", simp)
-apply (simp add: zmod_zadd1_eq)
-done
-
-lemma zmod_zadd_self2[simp]: "(b+a) mod a = b mod (a::int)"
-apply (case_tac "a = 0", simp)
-apply (simp add: zmod_zadd1_eq)
-done
-
-
-subsection{*Proving  @{term "a div (b*c) = (a div b) div c"} *}
-
-(*The condition c>0 seems necessary.  Consider that 7 div ~6 = ~2 but
-  7 div 2 div ~3 = 3 div ~3 = ~1.  The subcase (a div b) mod c = 0 seems
-  to cause particular problems.*)
-
-text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
-
-lemma zmult2_lemma_aux1: "[| (0::int) < c;  b < r;  r \<le> 0 |] ==> b*c < b*(q mod c) + r"
-apply (subgoal_tac "b * (c - q mod c) < r * 1")
-apply (simp add: right_diff_distrib)
-apply (rule order_le_less_trans)
-apply (erule_tac [2] mult_strict_right_mono)
-apply (rule mult_left_mono_neg)
-apply (auto simp add: compare_rls add_commute [of 1]
-                      add1_zle_eq pos_mod_bound)
-done
-
-lemma zmult2_lemma_aux2:
-     "[| (0::int) < c;   b < r;  r \<le> 0 |] ==> b * (q mod c) + r \<le> 0"
-apply (subgoal_tac "b * (q mod c) \<le> 0")
- apply arith
-apply (simp add: mult_le_0_iff)
-done
-
-lemma zmult2_lemma_aux3: "[| (0::int) < c;  0 \<le> r;  r < b |] ==> 0 \<le> b * (q mod c) + r"
-apply (subgoal_tac "0 \<le> b * (q mod c) ")
-apply arith
-apply (simp add: zero_le_mult_iff)
-done
-
-lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c"
-apply (subgoal_tac "r * 1 < b * (c - q mod c) ")
-apply (simp add: right_diff_distrib)
-apply (rule order_less_le_trans)
-apply (erule mult_strict_right_mono)
-apply (rule_tac [2] mult_left_mono)
-apply (auto simp add: compare_rls add_commute [of 1]
-                      add1_zle_eq pos_mod_bound)
-done
-
-lemma zmult2_lemma: "[| quorem ((a,b), (q,r));  b \<noteq> 0;  0 < c |]  
-      ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))"
-by (auto simp add: mult_ac quorem_def linorder_neq_iff
-                   zero_less_mult_iff right_distrib [symmetric] 
-                   zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4)
-
-lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c"
-apply (case_tac "b = 0", simp)
-apply (force simp add: quorem_div_mod [THEN zmult2_lemma, THEN quorem_div])
-done
-
-lemma zmod_zmult2_eq:
-     "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b"
-apply (case_tac "b = 0", simp)
-apply (force simp add: quorem_div_mod [THEN zmult2_lemma, THEN quorem_mod])
-done
-
-
-subsection{*Cancellation of Common Factors in div*}
-
-lemma zdiv_zmult_zmult1_aux1:
-     "[| (0::int) < b;  c \<noteq> 0 |] ==> (c*a) div (c*b) = a div b"
-by (subst zdiv_zmult2_eq, auto)
-
-lemma zdiv_zmult_zmult1_aux2:
-     "[| b < (0::int);  c \<noteq> 0 |] ==> (c*a) div (c*b) = a div b"
-apply (subgoal_tac " (c * (-a)) div (c * (-b)) = (-a) div (-b) ")
-apply (rule_tac [2] zdiv_zmult_zmult1_aux1, auto)
-done
-
-lemma zdiv_zmult_zmult1: "c \<noteq> (0::int) ==> (c*a) div (c*b) = a div b"
-apply (case_tac "b = 0", simp)
-apply (auto simp add: linorder_neq_iff zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2)
-done
-
-lemma zdiv_zmult_zmult2: "c \<noteq> (0::int) ==> (a*c) div (b*c) = a div b"
-apply (drule zdiv_zmult_zmult1)
-apply (auto simp add: mult_commute)
-done
-
-
-
-subsection{*Distribution of Factors over mod*}
-
-lemma zmod_zmult_zmult1_aux1:
-     "[| (0::int) < b;  c \<noteq> 0 |] ==> (c*a) mod (c*b) = c * (a mod b)"
-by (subst zmod_zmult2_eq, auto)
-
-lemma zmod_zmult_zmult1_aux2:
-     "[| b < (0::int);  c \<noteq> 0 |] ==> (c*a) mod (c*b) = c * (a mod b)"
-apply (subgoal_tac " (c * (-a)) mod (c * (-b)) = c * ((-a) mod (-b))")
-apply (rule_tac [2] zmod_zmult_zmult1_aux1, auto)
-done
-
-lemma zmod_zmult_zmult1: "(c*a) mod (c*b) = (c::int) * (a mod b)"
-apply (case_tac "b = 0", simp)
-apply (case_tac "c = 0", simp)
-apply (auto simp add: linorder_neq_iff zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2)
-done
-
-lemma zmod_zmult_zmult2: "(a*c) mod (b*c) = (a mod b) * (c::int)"
-apply (cut_tac c = c in zmod_zmult_zmult1)
-apply (auto simp add: mult_commute)
-done
-
-
-subsection {*Splitting Rules for div and mod*}
-
-text{*The proofs of the two lemmas below are essentially identical*}
-
-lemma split_pos_lemma:
- "0<k ==> 
-    P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)"
-apply (rule iffI, clarify)
- apply (erule_tac P="P ?x ?y" in rev_mp)  
- apply (subst zmod_zadd1_eq) 
- apply (subst zdiv_zadd1_eq) 
- apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial)  
-txt{*converse direction*}
-apply (drule_tac x = "n div k" in spec) 
-apply (drule_tac x = "n mod k" in spec, simp)
-done
-
-lemma split_neg_lemma:
- "k<0 ==>
-    P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)"
-apply (rule iffI, clarify)
- apply (erule_tac P="P ?x ?y" in rev_mp)  
- apply (subst zmod_zadd1_eq) 
- apply (subst zdiv_zadd1_eq) 
- apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial)  
-txt{*converse direction*}
-apply (drule_tac x = "n div k" in spec) 
-apply (drule_tac x = "n mod k" in spec, simp)
-done
-
-lemma split_zdiv:
- "P(n div k :: int) =
-  ((k = 0 --> P 0) & 
-   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & 
-   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))"
-apply (case_tac "k=0", simp)
-apply (simp only: linorder_neq_iff)
-apply (erule disjE) 
- apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] 
-                      split_neg_lemma [of concl: "%x y. P x"])
-done
-
-lemma split_zmod:
- "P(n mod k :: int) =
-  ((k = 0 --> P n) & 
-   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & 
-   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))"
-apply (case_tac "k=0", simp)
-apply (simp only: linorder_neq_iff)
-apply (erule disjE) 
- apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] 
-                      split_neg_lemma [of concl: "%x y. P y"])
-done
-
-(* Enable arith to deal with div 2 and mod 2: *)
-declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split]
-declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split]
-
-
-subsection{*Speeding up the Division Algorithm with Shifting*}
-
-text{*computing div by shifting *}
-
-lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a"
-proof cases
-  assume "a=0"
-    thus ?thesis by simp
-next
-  assume "a\<noteq>0" and le_a: "0\<le>a"   
-  hence a_pos: "1 \<le> a" by arith
-  hence one_less_a2: "1 < 2*a" by arith
-  hence le_2a: "2 * (1 + b mod a) \<le> 2 * a"
-    by (simp add: mult_le_cancel_left add_commute [of 1] add1_zle_eq)
-  with a_pos have "0 \<le> b mod a" by simp
-  hence le_addm: "0 \<le> 1 mod (2*a) + 2*(b mod a)"
-    by (simp add: mod_pos_pos_trivial one_less_a2)
-  with  le_2a
-  have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0"
-    by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2
-                  right_distrib) 
-  thus ?thesis
-    by (subst zdiv_zadd1_eq,
-        simp add: zdiv_zmult_zmult1 zmod_zmult_zmult1 one_less_a2
-                  div_pos_pos_trivial)
-qed
-
-lemma neg_zdiv_mult_2: "a \<le> (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a"
-apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ")
-apply (rule_tac [2] pos_zdiv_mult_2)
-apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib)
-apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))")
-apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric],
-       simp) 
-done
-
-
-(*Not clear why this must be proved separately; probably number_of causes
-  simplification problems*)
-lemma not_0_le_lemma: "~ 0 \<le> x ==> x \<le> (0::int)"
-by auto
-
-lemma zdiv_number_of_BIT[simp]:
-     "number_of (v BIT b) div number_of (w BIT bit.B0) =  
-          (if b=bit.B0 | (0::int) \<le> number_of w                    
-           then number_of v div (number_of w)     
-           else (number_of v + (1::int)) div (number_of w))"
-apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) 
-apply (simp add: zdiv_zmult_zmult1 pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac 
-            split: bit.split)
-done
-
-
-subsection{*Computing mod by Shifting (proofs resemble those for div)*}
-
-lemma pos_zmod_mult_2:
-     "(0::int) \<le> a ==> (1 + 2*b) mod (2*a) = 1 + 2 * (b mod a)"
-apply (case_tac "a = 0", simp)
-apply (subgoal_tac "1 < a * 2")
- prefer 2 apply arith
-apply (subgoal_tac "2* (1 + b mod a) \<le> 2*a")
- apply (rule_tac [2] mult_left_mono)
-apply (auto simp add: add_commute [of 1] mult_commute add1_zle_eq 
-                      pos_mod_bound)
-apply (subst zmod_zadd1_eq)
-apply (simp add: zmod_zmult_zmult2 mod_pos_pos_trivial)
-apply (rule mod_pos_pos_trivial)
-apply (auto simp add: mod_pos_pos_trivial left_distrib)
-apply (subgoal_tac "0 \<le> b mod a", arith, simp)
-done
-
-lemma neg_zmod_mult_2:
-     "a \<le> (0::int) ==> (1 + 2*b) mod (2*a) = 2 * ((b+1) mod a) - 1"
-apply (subgoal_tac "(1 + 2* (-b - 1)) mod (2* (-a)) = 
-                    1 + 2* ((-b - 1) mod (-a))")
-apply (rule_tac [2] pos_zmod_mult_2)
-apply (auto simp add: minus_mult_right [symmetric] right_diff_distrib)
-apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))")
- prefer 2 apply simp 
-apply (simp only: zmod_zminus_zminus diff_minus minus_add_distrib [symmetric])
-done
-
-lemma zmod_number_of_BIT [simp]:
-     "number_of (v BIT b) mod number_of (w BIT bit.B0) =  
-      (case b of
-          bit.B0 => 2 * (number_of v mod number_of w)
-        | bit.B1 => if (0::int) \<le> number_of w  
-                then 2 * (number_of v mod number_of w) + 1     
-                else 2 * ((number_of v + (1::int)) mod number_of w) - 1)"
-apply (simp only: number_of_eq numeral_simps UNIV_I split: bit.split) 
-apply (simp add: zmod_zmult_zmult1 pos_zmod_mult_2 
-                 not_0_le_lemma neg_zmod_mult_2 add_ac)
-done
-
-
-subsection{*Quotients of Signs*}
-
-lemma div_neg_pos_less0: "[| a < (0::int);  0 < b |] ==> a div b < 0"
-apply (subgoal_tac "a div b \<le> -1", force)
-apply (rule order_trans)
-apply (rule_tac a' = "-1" in zdiv_mono1)
-apply (auto simp add: zdiv_minus1)
-done
-
-lemma div_nonneg_neg_le0: "[| (0::int) \<le> a;  b < 0 |] ==> a div b \<le> 0"
-by (drule zdiv_mono1_neg, auto)
-
-lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)"
-apply auto
-apply (drule_tac [2] zdiv_mono1)
-apply (auto simp add: linorder_neq_iff)
-apply (simp (no_asm_use) add: linorder_not_less [symmetric])
-apply (blast intro: div_neg_pos_less0)
-done
-
-lemma neg_imp_zdiv_nonneg_iff:
-     "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))"
-apply (subst zdiv_zminus_zminus [symmetric])
-apply (subst pos_imp_zdiv_nonneg_iff, auto)
-done
-
-(*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*)
-lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)"
-by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff)
-
-(*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*)
-lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)"
-by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff)
-
-
-subsection {* The Divides Relation *}
-
-lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))"
-by(simp add:dvd_def zmod_eq_0_iff)
-
-lemmas zdvd_iff_zmod_eq_0_number_of [simp] =
-  zdvd_iff_zmod_eq_0 [of "number_of x" "number_of y", standard]
-
-lemma zdvd_0_right [iff]: "(m::int) dvd 0"
-by (simp add: dvd_def)
-
-lemma zdvd_0_left [iff]: "(0 dvd (m::int)) = (m = 0)"
-  by (simp add: dvd_def)
-
-lemma zdvd_1_left [iff]: "1 dvd (m::int)"
-  by (simp add: dvd_def)
-
-lemma zdvd_refl [simp]: "m dvd (m::int)"
-by (auto simp add: dvd_def intro: zmult_1_right [symmetric])
-
-lemma zdvd_trans: "m dvd n ==> n dvd k ==> m dvd (k::int)"
-by (auto simp add: dvd_def intro: mult_assoc)
-
-lemma zdvd_zminus_iff: "(m dvd -n) = (m dvd (n::int))"
-  apply (simp add: dvd_def, auto)
-   apply (rule_tac [!] x = "-k" in exI, auto)
-  done
-
-lemma zdvd_zminus2_iff: "(-m dvd n) = (m dvd (n::int))"
-  apply (simp add: dvd_def, auto)
-   apply (rule_tac [!] x = "-k" in exI, auto)
-  done
-lemma zdvd_abs1: "( \<bar>i::int\<bar> dvd j) = (i dvd j)" 
-  apply (cases "i > 0", simp)
-  apply (simp add: dvd_def)
-  apply (rule iffI)
-  apply (erule exE)
-  apply (rule_tac x="- k" in exI, simp)
-  apply (erule exE)
-  apply (rule_tac x="- k" in exI, simp)
-  done
-lemma zdvd_abs2: "( (i::int) dvd \<bar>j\<bar>) = (i dvd j)" 
-  apply (cases "j > 0", simp)
-  apply (simp add: dvd_def)
-  apply (rule iffI)
-  apply (erule exE)
-  apply (rule_tac x="- k" in exI, simp)
-  apply (erule exE)
-  apply (rule_tac x="- k" in exI, simp)
-  done
-
-lemma zdvd_anti_sym:
-    "0 < m ==> 0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)"
-  apply (simp add: dvd_def, auto)
-  apply (simp add: mult_assoc zero_less_mult_iff zmult_eq_1_iff)
-  done
-
-lemma zdvd_zadd: "k dvd m ==> k dvd n ==> k dvd (m + n :: int)"
-  apply (simp add: dvd_def)
-  apply (blast intro: right_distrib [symmetric])
-  done
-
-lemma zdvd_dvd_eq: assumes anz:"a \<noteq> 0" and ab: "(a::int) dvd b" and ba:"b dvd a" 
-  shows "\<bar>a\<bar> = \<bar>b\<bar>"
-proof-
-  from ab obtain k where k:"b = a*k" unfolding dvd_def by blast 
-  from ba obtain k' where k':"a = b*k'" unfolding dvd_def by blast 
-  from k k' have "a = a*k*k'" by simp
-  with mult_cancel_left1[where c="a" and b="k*k'"]
-  have kk':"k*k' = 1" using anz by (simp add: mult_assoc)
-  hence "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1" by (simp add: zmult_eq_1_iff)
-  thus ?thesis using k k' by auto
-qed
-
-lemma zdvd_zdiff: "k dvd m ==> k dvd n ==> k dvd (m - n :: int)"
-  apply (simp add: dvd_def)
-  apply (blast intro: right_diff_distrib [symmetric])
-  done
-
-lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)"
-  apply (subgoal_tac "m = n + (m - n)")
-   apply (erule ssubst)
-   apply (blast intro: zdvd_zadd, simp)
-  done
-
-lemma zdvd_zmult: "k dvd (n::int) ==> k dvd m * n"
-  apply (simp add: dvd_def)
-  apply (blast intro: mult_left_commute)
-  done
-
-lemma zdvd_zmult2: "k dvd (m::int) ==> k dvd m * n"
-  apply (subst mult_commute)
-  apply (erule zdvd_zmult)
-  done
-
-lemma zdvd_triv_right [iff]: "(k::int) dvd m * k"
-  apply (rule zdvd_zmult)
-  apply (rule zdvd_refl)
-  done
-
-lemma zdvd_triv_left [iff]: "(k::int) dvd k * m"
-  apply (rule zdvd_zmult2)
-  apply (rule zdvd_refl)
-  done
-
-lemma zdvd_zmultD2: "j * k dvd n ==> j dvd (n::int)"
-  apply (simp add: dvd_def)
-  apply (simp add: mult_assoc, blast)
-  done
-
-lemma zdvd_zmultD: "j * k dvd n ==> k dvd (n::int)"
-  apply (rule zdvd_zmultD2)
-  apply (subst mult_commute, assumption)
-  done
-
-lemma zdvd_zmult_mono: "i dvd m ==> j dvd (n::int) ==> i * j dvd m * n"
-  apply (simp add: dvd_def, clarify)
-  apply (rule_tac x = "k * ka" in exI)
-  apply (simp add: mult_ac)
-  done
-
-lemma zdvd_reduce: "(k dvd n + k * m) = (k dvd (n::int))"
-  apply (rule iffI)
-   apply (erule_tac [2] zdvd_zadd)
-   apply (subgoal_tac "n = (n + k * m) - k * m")
-    apply (erule ssubst)
-    apply (erule zdvd_zdiff, simp_all)
-  done
-
-lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n"
-  apply (simp add: dvd_def)
-  apply (auto simp add: zmod_zmult_zmult1)
-  done
-
-lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)"
-  apply (subgoal_tac "k dvd n * (m div n) + m mod n")
-   apply (simp add: zmod_zdiv_equality [symmetric])
-  apply (simp only: zdvd_zadd zdvd_zmult2)
-  done
-
-lemma zdvd_not_zless: "0 < m ==> m < n ==> \<not> n dvd (m::int)"
-  apply (simp add: dvd_def, auto)
-  apply (subgoal_tac "0 < n")
-   prefer 2
-   apply (blast intro: order_less_trans)
-  apply (simp add: zero_less_mult_iff)
-  apply (subgoal_tac "n * k < n * 1")
-   apply (drule mult_less_cancel_left [THEN iffD1], auto)
-  done
-lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)"
-  using zmod_zdiv_equality[where a="m" and b="n"]
-  by (simp add: ring_eq_simps)
-
-lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m"
-apply (subgoal_tac "m mod n = 0")
- apply (simp add: zmult_div_cancel)
-apply (simp only: zdvd_iff_zmod_eq_0)
-done
-
-lemma zdvd_mult_cancel: assumes d:"k * m dvd k * n" and kz:"k \<noteq> (0::int)"
-  shows "m dvd n"
-proof-
-  from d obtain h where h: "k*n = k*m * h" unfolding dvd_def by blast
-  {assume "n \<noteq> m*h" hence "k* n \<noteq> k* (m*h)" using kz by simp
-    with h have False by (simp add: mult_assoc)}
-  hence "n = m * h" by blast
-  thus ?thesis by blast
-qed
-
-theorem ex_nat: "(\<exists>x::nat. P x) = (\<exists>x::int. 0 <= x \<and> P (nat x))"
-  apply (simp split add: split_nat)
-  apply (rule iffI)
-  apply (erule exE)
-  apply (rule_tac x = "int x" in exI)
-  apply simp
-  apply (erule exE)
-  apply (rule_tac x = "nat x" in exI)
-  apply (erule conjE)
-  apply (erule_tac x = "nat x" in allE)
-  apply simp
-  done
-
-theorem zdvd_int: "(x dvd y) = (int x dvd int y)"
-  apply (simp only: dvd_def ex_nat int_int_eq [symmetric] zmult_int [symmetric]
-    nat_0_le cong add: conj_cong)
-  apply (rule iffI)
-  apply iprover
-  apply (erule exE)
-  apply (case_tac "x=0")
-  apply (rule_tac x=0 in exI)
-  apply simp
-  apply (case_tac "0 \<le> k")
-  apply iprover
-  apply (simp add: linorder_not_le)
-  apply (drule mult_strict_left_mono_neg [OF iffD2 [OF zero_less_int_conv]])
-  apply assumption
-  apply (simp add: mult_ac)
-  done
-
-lemma zdvd1_eq[simp]: "(x::int) dvd 1 = ( \<bar>x\<bar> = 1)"
-proof
-  assume d: "x dvd 1" hence "int (nat \<bar>x\<bar>) dvd int (nat 1)" by (simp add: zdvd_abs1)
-  hence "nat \<bar>x\<bar> dvd 1" by (simp add: zdvd_int)
-  hence "nat \<bar>x\<bar> = 1"  by simp
-  thus "\<bar>x\<bar> = 1" by (cases "x < 0", auto)
-next
-  assume "\<bar>x\<bar>=1" thus "x dvd 1" 
-    by(cases "x < 0",simp_all add: minus_equation_iff zdvd_iff_zmod_eq_0)
-qed
-lemma zdvd_mult_cancel1: 
-  assumes mp:"m \<noteq>(0::int)" shows "(m * n dvd m) = (\<bar>n\<bar> = 1)"
-proof
-  assume n1: "\<bar>n\<bar> = 1" thus "m * n dvd m" 
-    by (cases "n >0", auto simp add: zdvd_zminus2_iff minus_equation_iff)
-next
-  assume H: "m * n dvd m" hence H2: "m * n dvd m * 1" by simp
-  from zdvd_mult_cancel[OF H2 mp] show "\<bar>n\<bar> = 1" by (simp only: zdvd1_eq)
-qed
-
-lemma int_dvd_iff: "(int m dvd z) = (m dvd nat (abs z))"
-  apply (auto simp add: dvd_def nat_abs_mult_distrib)
-  apply (auto simp add: nat_eq_iff abs_if split add: split_if_asm)
-   apply (rule_tac x = "-(int k)" in exI)
-  apply (auto simp add: int_mult)
-  done
-
-lemma dvd_int_iff: "(z dvd int m) = (nat (abs z) dvd m)"
-  apply (auto simp add: dvd_def abs_if int_mult)
-    apply (rule_tac [3] x = "nat k" in exI)
-    apply (rule_tac [2] x = "-(int k)" in exI)
-    apply (rule_tac x = "nat (-k)" in exI)
-    apply (cut_tac [3] k = m in int_less_0_conv)
-    apply (cut_tac k = m in int_less_0_conv)
-    apply (auto simp add: zero_le_mult_iff mult_less_0_iff
-      nat_mult_distrib [symmetric] nat_eq_iff2)
-  done
-
-lemma nat_dvd_iff: "(nat z dvd m) = (if 0 \<le> z then (z dvd int m) else m = 0)"
-  apply (auto simp add: dvd_def int_mult)
-  apply (rule_tac x = "nat k" in exI)
-  apply (cut_tac k = m in int_less_0_conv)
-  apply (auto simp add: zero_le_mult_iff mult_less_0_iff
-    nat_mult_distrib [symmetric] nat_eq_iff2)
-  done
-
-lemma zminus_dvd_iff [iff]: "(-z dvd w) = (z dvd (w::int))"
-  apply (auto simp add: dvd_def)
-   apply (rule_tac [!] x = "-k" in exI, auto)
-  done
-
-lemma dvd_zminus_iff [iff]: "(z dvd -w) = (z dvd (w::int))"
-  apply (auto simp add: dvd_def)
-   apply (drule minus_equation_iff [THEN iffD1])
-   apply (rule_tac [!] x = "-k" in exI, auto)
-  done
-
-lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \<le> (n::int)"
-  apply (rule_tac z=n in int_cases)
-  apply (auto simp add: dvd_int_iff) 
-  apply (rule_tac z=z in int_cases) 
-  apply (auto simp add: dvd_imp_le) 
-  done
-
-
-subsection{*Integer Powers*} 
-
-instance int :: power ..
-
-primrec
-  "p ^ 0 = 1"
-  "p ^ (Suc n) = (p::int) * (p ^ n)"
-
-
-instance int :: recpower
-proof
-  fix z :: int
-  fix n :: nat
-  show "z^0 = 1" by simp
-  show "z^(Suc n) = z * (z^n)" by simp
-qed
-
-
-lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m"
-apply (induct "y", auto)
-apply (rule zmod_zmult1_eq [THEN trans])
-apply (simp (no_asm_simp))
-apply (rule zmod_zmult_distrib [symmetric])
-done
-
-lemma zpower_zadd_distrib: "x^(y+z) = ((x^y)*(x^z)::int)"
-  by (rule Power.power_add)
-
-lemma zpower_zpower: "(x^y)^z = (x^(y*z)::int)"
-  by (rule Power.power_mult [symmetric])
-
-lemma zero_less_zpower_abs_iff [simp]:
-     "(0 < (abs x)^n) = (x \<noteq> (0::int) | n=0)"
-apply (induct "n")
-apply (auto simp add: zero_less_mult_iff)
-done
-
-lemma zero_le_zpower_abs [simp]: "(0::int) <= (abs x)^n"
-apply (induct "n")
-apply (auto simp add: zero_le_mult_iff)
-done
-
-lemma int_power: "int (m^n) = (int m) ^ n"
-  by (induct n, simp_all add: int_mult)
-
-text{*Compatibility binding*}
-lemmas zpower_int = int_power [symmetric]
-
-lemma zdiv_int: "int (a div b) = (int a) div (int b)"
-apply (subst split_div, auto)
-apply (subst split_zdiv, auto)
-apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient)
-apply (auto simp add: IntDiv.quorem_def int_eq_of_nat)
-done
-
-lemma zmod_int: "int (a mod b) = (int a) mod (int b)"
-apply (subst split_mod, auto)
-apply (subst split_zmod, auto)
-apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia 
-       in unique_remainder)
-apply (auto simp add: IntDiv.quorem_def int_eq_of_nat)
-done
-
-text{*Suggested by Matthias Daum*}
-lemma int_power_div_base:
-     "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"
-apply (subgoal_tac "k ^ m = k ^ ((m - 1) + 1)")
- apply (erule ssubst)
- apply (simp only: power_add)
- apply simp_all
-done
-
-text {* code serializer setup *}
-
-code_modulename SML
-  IntDiv Integer
-
-code_modulename OCaml
-  IntDiv Integer
-
-code_modulename Haskell
-  IntDiv Divides
-
-end
--- a/src/HOL/Integ/NatBin.thy	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,903 +0,0 @@
-(*  Title:      HOL/NatBin.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1999  University of Cambridge
-*)
-
-header {* Binary arithmetic for the natural numbers *}
-
-theory NatBin
-imports IntDiv
-begin
-
-text {*
-  Arithmetic for naturals is reduced to that for the non-negative integers.
-*}
-
-instance nat :: number
-  nat_number_of_def [code inline]: "number_of v == nat (number_of (v\<Colon>int))" ..
-
-abbreviation (xsymbols)
-  square :: "'a::power => 'a"  ("(_\<twosuperior>)" [1000] 999) where
-  "x\<twosuperior> == x^2"
-
-notation (latex output)
-  square  ("(_\<twosuperior>)" [1000] 999)
-
-notation (HTML output)
-  square  ("(_\<twosuperior>)" [1000] 999)
-
-
-subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*}
-
-declare nat_0 [simp] nat_1 [simp]
-
-lemma nat_number_of [simp]: "nat (number_of w) = number_of w"
-by (simp add: nat_number_of_def)
-
-lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)"
-by (simp add: nat_number_of_def)
-
-lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
-by (simp add: nat_1 nat_number_of_def)
-
-lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
-by (simp add: nat_numeral_1_eq_1)
-
-lemma numeral_2_eq_2: "2 = Suc (Suc 0)"
-apply (unfold nat_number_of_def)
-apply (rule nat_2)
-done
-
-
-text{*Distributive laws for type @{text nat}.  The others are in theory
-   @{text IntArith}, but these require div and mod to be defined for type
-   "int".  They also need some of the lemmas proved above.*}
-
-lemma nat_div_distrib: "(0::int) <= z ==> nat (z div z') = nat z div nat z'"
-apply (case_tac "0 <= z'")
-apply (auto simp add: div_nonneg_neg_le0 DIVISION_BY_ZERO_DIV)
-apply (case_tac "z' = 0", simp add: DIVISION_BY_ZERO)
-apply (auto elim!: nonneg_eq_int)
-apply (rename_tac m m')
-apply (subgoal_tac "0 <= int m div int m'")
- prefer 2 apply (simp add: nat_numeral_0_eq_0 pos_imp_zdiv_nonneg_iff) 
-apply (rule inj_int [THEN injD], simp)
-apply (rule_tac r = "int (m mod m') " in quorem_div)
- prefer 2 apply force
-apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int 
-                 zmult_int)
-done
-
-(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*)
-lemma nat_mod_distrib:
-     "[| (0::int) <= z;  0 <= z' |] ==> nat (z mod z') = nat z mod nat z'"
-apply (case_tac "z' = 0", simp add: DIVISION_BY_ZERO)
-apply (auto elim!: nonneg_eq_int)
-apply (rename_tac m m')
-apply (subgoal_tac "0 <= int m mod int m'")
- prefer 2 apply (simp add: nat_less_iff nat_numeral_0_eq_0 pos_mod_sign) 
-apply (rule inj_int [THEN injD], simp)
-apply (rule_tac q = "int (m div m') " in quorem_mod)
- prefer 2 apply force
-apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int zmult_int)
-done
-
-text{*Suggested by Matthias Daum*}
-lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)"
-apply (subgoal_tac "nat x div nat k < nat x")
- apply (simp (asm_lr) add: nat_div_distrib [symmetric])
-apply (rule Divides.div_less_dividend, simp_all) 
-done
-
-subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-lemma int_nat_number_of [simp]:
-     "int (number_of v :: nat) =  
-         (if neg (number_of v :: int) then 0  
-          else (number_of v :: int))"
-by (simp del: nat_number_of
-	 add: neg_nat nat_number_of_def not_neg_nat add_assoc)
-
-
-subsubsection{*Successor *}
-
-lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)"
-apply (rule sym)
-apply (simp add: nat_eq_iff int_Suc)
-done
-
-lemma Suc_nat_number_of_add:
-     "Suc (number_of v + n) =  
-        (if neg (number_of v :: int) then 1+n else number_of (Numeral.succ v) + n)" 
-by (simp del: nat_number_of 
-         add: nat_number_of_def neg_nat
-              Suc_nat_eq_nat_zadd1 number_of_succ) 
-
-lemma Suc_nat_number_of [simp]:
-     "Suc (number_of v) =  
-        (if neg (number_of v :: int) then 1 else number_of (Numeral.succ v))"
-apply (cut_tac n = 0 in Suc_nat_number_of_add)
-apply (simp cong del: if_weak_cong)
-done
-
-
-subsubsection{*Addition *}
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-lemma add_nat_number_of [simp]:
-     "(number_of v :: nat) + number_of v' =  
-         (if neg (number_of v :: int) then number_of v'  
-          else if neg (number_of v' :: int) then number_of v  
-          else number_of (v + v'))"
-by (force dest!: neg_nat
-          simp del: nat_number_of
-          simp add: nat_number_of_def nat_add_distrib [symmetric]) 
-
-
-subsubsection{*Subtraction *}
-
-lemma diff_nat_eq_if:
-     "nat z - nat z' =  
-        (if neg z' then nat z   
-         else let d = z-z' in     
-              if neg d then 0 else nat d)"
-apply (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0)
-done
-
-lemma diff_nat_number_of [simp]: 
-     "(number_of v :: nat) - number_of v' =  
-        (if neg (number_of v' :: int) then number_of v  
-         else let d = number_of (v + uminus v') in     
-              if neg d then 0 else nat d)"
-by (simp del: nat_number_of add: diff_nat_eq_if nat_number_of_def) 
-
-
-
-subsubsection{*Multiplication *}
-
-lemma mult_nat_number_of [simp]:
-     "(number_of v :: nat) * number_of v' =  
-       (if neg (number_of v :: int) then 0 else number_of (v * v'))"
-by (force dest!: neg_nat
-          simp del: nat_number_of
-          simp add: nat_number_of_def nat_mult_distrib [symmetric]) 
-
-
-
-subsubsection{*Quotient *}
-
-lemma div_nat_number_of [simp]:
-     "(number_of v :: nat)  div  number_of v' =  
-          (if neg (number_of v :: int) then 0  
-           else nat (number_of v div number_of v'))"
-by (force dest!: neg_nat
-          simp del: nat_number_of
-          simp add: nat_number_of_def nat_div_distrib [symmetric]) 
-
-lemma one_div_nat_number_of [simp]:
-     "(Suc 0)  div  number_of v' = (nat (1 div number_of v'))" 
-by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
-
-
-subsubsection{*Remainder *}
-
-lemma mod_nat_number_of [simp]:
-     "(number_of v :: nat)  mod  number_of v' =  
-        (if neg (number_of v :: int) then 0  
-         else if neg (number_of v' :: int) then number_of v  
-         else nat (number_of v mod number_of v'))"
-by (force dest!: neg_nat
-          simp del: nat_number_of
-          simp add: nat_number_of_def nat_mod_distrib [symmetric]) 
-
-lemma one_mod_nat_number_of [simp]:
-     "(Suc 0)  mod  number_of v' =  
-        (if neg (number_of v' :: int) then Suc 0
-         else nat (1 mod number_of v'))"
-by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
-
-
-subsubsection{* Divisibility *}
-
-lemmas dvd_eq_mod_eq_0_number_of =
-  dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard]
-
-declare dvd_eq_mod_eq_0_number_of [simp]
-
-ML
-{*
-val nat_number_of_def = thm"nat_number_of_def";
-
-val nat_number_of = thm"nat_number_of";
-val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0";
-val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1";
-val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0";
-val numeral_2_eq_2 = thm"numeral_2_eq_2";
-val nat_div_distrib = thm"nat_div_distrib";
-val nat_mod_distrib = thm"nat_mod_distrib";
-val int_nat_number_of = thm"int_nat_number_of";
-val Suc_nat_eq_nat_zadd1 = thm"Suc_nat_eq_nat_zadd1";
-val Suc_nat_number_of_add = thm"Suc_nat_number_of_add";
-val Suc_nat_number_of = thm"Suc_nat_number_of";
-val add_nat_number_of = thm"add_nat_number_of";
-val diff_nat_eq_if = thm"diff_nat_eq_if";
-val diff_nat_number_of = thm"diff_nat_number_of";
-val mult_nat_number_of = thm"mult_nat_number_of";
-val div_nat_number_of = thm"div_nat_number_of";
-val mod_nat_number_of = thm"mod_nat_number_of";
-*}
-
-
-subsection{*Comparisons*}
-
-subsubsection{*Equals (=) *}
-
-lemma eq_nat_nat_iff:
-     "[| (0::int) <= z;  0 <= z' |] ==> (nat z = nat z') = (z=z')"
-by (auto elim!: nonneg_eq_int)
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-lemma eq_nat_number_of [simp]:
-     "((number_of v :: nat) = number_of v') =  
-      (if neg (number_of v :: int) then (iszero (number_of v' :: int) | neg (number_of v' :: int))  
-       else if neg (number_of v' :: int) then iszero (number_of v :: int)  
-       else iszero (number_of (v + uminus v') :: int))"
-apply (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
-                  eq_nat_nat_iff eq_number_of_eq nat_0 iszero_def
-            split add: split_if cong add: imp_cong)
-apply (simp only: nat_eq_iff nat_eq_iff2)
-apply (simp add: not_neg_eq_ge_0 [symmetric])
-done
-
-
-subsubsection{*Less-than (<) *}
-
-(*"neg" is used in rewrite rules for binary comparisons*)
-lemma less_nat_number_of [simp]:
-     "((number_of v :: nat) < number_of v') =  
-         (if neg (number_of v :: int) then neg (number_of (uminus v') :: int)  
-          else neg (number_of (v + uminus v') :: int))"
-by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
-                nat_less_eq_zless less_number_of_eq_neg zless_nat_eq_int_zless
-         cong add: imp_cong, simp add: Pls_def)
-
-
-(*Maps #n to n for n = 0, 1, 2*)
-lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2
-
-
-subsection{*Powers with Numeric Exponents*}
-
-text{*We cannot refer to the number @{term 2} in @{text Ring_and_Field.thy}.
-We cannot prove general results about the numeral @{term "-1"}, so we have to
-use @{term "- 1"} instead.*}
-
-lemma power2_eq_square: "(a::'a::{comm_semiring_1_cancel,recpower})\<twosuperior> = a * a"
-  by (simp add: numeral_2_eq_2 Power.power_Suc)
-
-lemma zero_power2 [simp]: "(0::'a::{comm_semiring_1_cancel,recpower})\<twosuperior> = 0"
-  by (simp add: power2_eq_square)
-
-lemma one_power2 [simp]: "(1::'a::{comm_semiring_1_cancel,recpower})\<twosuperior> = 1"
-  by (simp add: power2_eq_square)
-
-lemma power3_eq_cube: "(x::'a::recpower) ^ 3 = x * x * x"
-  apply (subgoal_tac "3 = Suc (Suc (Suc 0))")
-  apply (erule ssubst)
-  apply (simp add: power_Suc mult_ac)
-  apply (unfold nat_number_of_def)
-  apply (subst nat_eq_iff)
-  apply simp
-done
-
-text{*Squares of literal numerals will be evaluated.*}
-lemmas power2_eq_square_number_of =
-    power2_eq_square [of "number_of w", standard]
-declare power2_eq_square_number_of [simp]
-
-
-lemma zero_le_power2[simp]: "0 \<le> (a\<twosuperior>::'a::{ordered_idom,recpower})"
-  by (simp add: power2_eq_square)
-
-lemma zero_less_power2[simp]:
-     "(0 < a\<twosuperior>) = (a \<noteq> (0::'a::{ordered_idom,recpower}))"
-  by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
-
-lemma power2_less_0[simp]:
-  fixes a :: "'a::{ordered_idom,recpower}"
-  shows "~ (a\<twosuperior> < 0)"
-by (force simp add: power2_eq_square mult_less_0_iff) 
-
-lemma zero_eq_power2[simp]:
-     "(a\<twosuperior> = 0) = (a = (0::'a::{ordered_idom,recpower}))"
-  by (force simp add: power2_eq_square mult_eq_0_iff)
-
-lemma abs_power2[simp]:
-     "abs(a\<twosuperior>) = (a\<twosuperior>::'a::{ordered_idom,recpower})"
-  by (simp add: power2_eq_square abs_mult abs_mult_self)
-
-lemma power2_abs[simp]:
-     "(abs a)\<twosuperior> = (a\<twosuperior>::'a::{ordered_idom,recpower})"
-  by (simp add: power2_eq_square abs_mult_self)
-
-lemma power2_minus[simp]:
-     "(- a)\<twosuperior> = (a\<twosuperior>::'a::{comm_ring_1,recpower})"
-  by (simp add: power2_eq_square)
-
-lemma power2_le_imp_le:
-  fixes x y :: "'a::{ordered_semidom,recpower}"
-  shows "\<lbrakk>x\<twosuperior> \<le> y\<twosuperior>; 0 \<le> y\<rbrakk> \<Longrightarrow> x \<le> y"
-unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)
-
-lemma power2_less_imp_less:
-  fixes x y :: "'a::{ordered_semidom,recpower}"
-  shows "\<lbrakk>x\<twosuperior> < y\<twosuperior>; 0 \<le> y\<rbrakk> \<Longrightarrow> x < y"
-by (rule power_less_imp_less_base)
-
-lemma power2_eq_imp_eq:
-  fixes x y :: "'a::{ordered_semidom,recpower}"
-  shows "\<lbrakk>x\<twosuperior> = y\<twosuperior>; 0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> x = y"
-unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base, simp)
-
-lemma power_minus1_even[simp]: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,recpower})"
-apply (induct "n")
-apply (auto simp add: power_Suc power_add)
-done
-
-lemma power_even_eq: "(a::'a::recpower) ^ (2*n) = (a^n)^2"
-by (subst mult_commute) (simp add: power_mult)
-
-lemma power_odd_eq: "(a::int) ^ Suc(2*n) = a * (a^n)^2"
-by (simp add: power_even_eq) 
-
-lemma power_minus_even [simp]:
-     "(-a) ^ (2*n) = (a::'a::{comm_ring_1,recpower}) ^ (2*n)"
-by (simp add: power_minus1_even power_minus [of a]) 
-
-lemma zero_le_even_power'[simp]:
-     "0 \<le> (a::'a::{ordered_idom,recpower}) ^ (2*n)"
-proof (induct "n")
-  case 0
-    show ?case by (simp add: zero_le_one)
-next
-  case (Suc n)
-    have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" 
-      by (simp add: mult_ac power_add power2_eq_square)
-    thus ?case
-      by (simp add: prems zero_le_mult_iff)
-qed
-
-lemma odd_power_less_zero:
-     "(a::'a::{ordered_idom,recpower}) < 0 ==> a ^ Suc(2*n) < 0"
-proof (induct "n")
-  case 0
-    show ?case by (simp add: Power.power_Suc)
-next
-  case (Suc n)
-    have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" 
-      by (simp add: mult_ac power_add power2_eq_square Power.power_Suc)
-    thus ?case
-      by (simp add: prems mult_less_0_iff mult_neg_neg)
-qed
-
-lemma odd_0_le_power_imp_0_le:
-     "0 \<le> a  ^ Suc(2*n) ==> 0 \<le> (a::'a::{ordered_idom,recpower})"
-apply (insert odd_power_less_zero [of a n]) 
-apply (force simp add: linorder_not_less [symmetric]) 
-done
-
-text{*Simprules for comparisons where common factors can be cancelled.*}
-lemmas zero_compare_simps =
-    add_strict_increasing add_strict_increasing2 add_increasing
-    zero_le_mult_iff zero_le_divide_iff 
-    zero_less_mult_iff zero_less_divide_iff 
-    mult_le_0_iff divide_le_0_iff 
-    mult_less_0_iff divide_less_0_iff 
-    zero_le_power2 power2_less_0
-
-subsubsection{*Nat *}
-
-lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
-by (simp add: numerals)
-
-(*Expresses a natural number constant as the Suc of another one.
-  NOT suitable for rewriting because n recurs in the condition.*)
-lemmas expand_Suc = Suc_pred' [of "number_of v", standard]
-
-subsubsection{*Arith *}
-
-lemma Suc_eq_add_numeral_1: "Suc n = n + 1"
-by (simp add: numerals)
-
-lemma Suc_eq_add_numeral_1_left: "Suc n = 1 + n"
-by (simp add: numerals)
-
-(* These two can be useful when m = number_of... *)
-
-lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))"
-apply (case_tac "m")
-apply (simp_all add: numerals)
-done
-
-lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))"
-apply (case_tac "m")
-apply (simp_all add: numerals)
-done
-
-lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))"
-apply (case_tac "m")
-apply (simp_all add: numerals)
-done
-
-
-subsection{*Comparisons involving (0::nat) *}
-
-text{*Simplification already does @{term "n<0"}, @{term "n\<le>0"} and @{term "0\<le>n"}.*}
-
-lemma eq_number_of_0 [simp]:
-     "(number_of v = (0::nat)) =  
-      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
-by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
-
-lemma eq_0_number_of [simp]:
-     "((0::nat) = number_of v) =  
-      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
-by (rule trans [OF eq_sym_conv eq_number_of_0])
-
-lemma less_0_number_of [simp]:
-     "((0::nat) < number_of v) = neg (number_of (uminus v) :: int)"
-by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] Pls_def)
-
-
-lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)"
-by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
-
-
-
-subsection{*Comparisons involving  @{term Suc} *}
-
-lemma eq_number_of_Suc [simp]:
-     "(number_of v = Suc n) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then False else nat pv = n)"
-apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
-                  number_of_pred nat_number_of_def 
-            split add: split_if)
-apply (rule_tac x = "number_of v" in spec)
-apply (auto simp add: nat_eq_iff)
-done
-
-lemma Suc_eq_number_of [simp]:
-     "(Suc n = number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then False else nat pv = n)"
-by (rule trans [OF eq_sym_conv eq_number_of_Suc])
-
-lemma less_number_of_Suc [simp]:
-     "(number_of v < Suc n) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then True else nat pv < n)"
-apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
-                  number_of_pred nat_number_of_def  
-            split add: split_if)
-apply (rule_tac x = "number_of v" in spec)
-apply (auto simp add: nat_less_iff)
-done
-
-lemma less_Suc_number_of [simp]:
-     "(Suc n < number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then False else n < nat pv)"
-apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
-                  number_of_pred nat_number_of_def
-            split add: split_if)
-apply (rule_tac x = "number_of v" in spec)
-apply (auto simp add: zless_nat_eq_int_zless)
-done
-
-lemma le_number_of_Suc [simp]:
-     "(number_of v <= Suc n) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then True else nat pv <= n)"
-by (simp add: Let_def less_Suc_number_of linorder_not_less [symmetric])
-
-lemma le_Suc_number_of [simp]:
-     "(Suc n <= number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then False else n <= nat pv)"
-by (simp add: Let_def less_number_of_Suc linorder_not_less [symmetric])
-
-
-(* Push int(.) inwards: *)
-declare zadd_int [symmetric, simp]
-
-lemma lemma1: "(m+m = n+n) = (m = (n::int))"
-by auto
-
-lemma lemma2: "m+m ~= (1::int) + (n + n)"
-apply auto
-apply (drule_tac f = "%x. x mod 2" in arg_cong)
-apply (simp add: zmod_zadd1_eq)
-done
-
-lemma eq_number_of_BIT_BIT:
-     "((number_of (v BIT x) ::int) = number_of (w BIT y)) =  
-      (x=y & (((number_of v) ::int) = number_of w))"
-apply (simp only: number_of_BIT lemma1 lemma2 eq_commute
-               OrderedGroup.add_left_cancel add_assoc OrderedGroup.add_0_left
-            split add: bit.split)
-apply simp
-done
-
-lemma eq_number_of_BIT_Pls:
-     "((number_of (v BIT x) ::int) = Numeral0) =  
-      (x=bit.B0 & (((number_of v) ::int) = Numeral0))"
-apply (simp only: simp_thms  add: number_of_BIT number_of_Pls eq_commute
-            split add: bit.split cong: imp_cong)
-apply (rule_tac x = "number_of v" in spec, safe)
-apply (simp_all (no_asm_use))
-apply (drule_tac f = "%x. x mod 2" in arg_cong)
-apply (simp add: zmod_zadd1_eq)
-done
-
-lemma eq_number_of_BIT_Min:
-     "((number_of (v BIT x) ::int) = number_of Numeral.Min) =  
-      (x=bit.B1 & (((number_of v) ::int) = number_of Numeral.Min))"
-apply (simp only: simp_thms  add: number_of_BIT number_of_Min eq_commute
-            split add: bit.split cong: imp_cong)
-apply (rule_tac x = "number_of v" in spec, auto)
-apply (drule_tac f = "%x. x mod 2" in arg_cong, auto)
-done
-
-lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Numeral.Min"
-by auto
-
-
-
-subsection{*Max and Min Combined with @{term Suc} *}
-
-lemma max_number_of_Suc [simp]:
-     "max (Suc n) (number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then Suc n else Suc(max n (nat pv)))"
-apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
-            split add: split_if nat.split)
-apply (rule_tac x = "number_of v" in spec) 
-apply auto
-done
- 
-lemma max_Suc_number_of [simp]:
-     "max (number_of v) (Suc n) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then Suc n else Suc(max (nat pv) n))"
-apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
-            split add: split_if nat.split)
-apply (rule_tac x = "number_of v" in spec) 
-apply auto
-done
- 
-lemma min_number_of_Suc [simp]:
-     "min (Suc n) (number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then 0 else Suc(min n (nat pv)))"
-apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
-            split add: split_if nat.split)
-apply (rule_tac x = "number_of v" in spec) 
-apply auto
-done
- 
-lemma min_Suc_number_of [simp]:
-     "min (number_of v) (Suc n) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then 0 else Suc(min (nat pv) n))"
-apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
-            split add: split_if nat.split)
-apply (rule_tac x = "number_of v" in spec) 
-apply auto
-done
- 
-subsection{*Literal arithmetic involving powers*}
-
-lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n"
-apply (induct "n")
-apply (simp_all (no_asm_simp) add: nat_mult_distrib)
-done
-
-lemma power_nat_number_of:
-     "(number_of v :: nat) ^ n =  
-       (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))"
-by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq
-         split add: split_if cong: imp_cong)
-
-
-lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w", standard]
-declare power_nat_number_of_number_of [simp]
-
-
-
-text{*For the integers*}
-
-lemma zpower_number_of_even:
-  "(z::int) ^ number_of (w BIT bit.B0) = (let w = z ^ (number_of w) in w * w)"
-unfolding Let_def nat_number_of_def number_of_BIT bit.cases
-apply (rule_tac x = "number_of w" in spec, clarify)
-apply (case_tac " (0::int) <= x")
-apply (auto simp add: nat_mult_distrib power_even_eq power2_eq_square)
-done
-
-lemma zpower_number_of_odd:
-  "(z::int) ^ number_of (w BIT bit.B1) = (if (0::int) <= number_of w                    
-     then (let w = z ^ (number_of w) in z * w * w) else 1)"
-unfolding Let_def nat_number_of_def number_of_BIT bit.cases
-apply (rule_tac x = "number_of w" in spec, auto)
-apply (simp only: nat_add_distrib nat_mult_distrib)
-apply simp
-apply (auto simp add: nat_add_distrib nat_mult_distrib power_even_eq power2_eq_square neg_nat)
-done
-
-lemmas zpower_number_of_even_number_of =
-    zpower_number_of_even [of "number_of v", standard]
-declare zpower_number_of_even_number_of [simp]
-
-lemmas zpower_number_of_odd_number_of =
-    zpower_number_of_odd [of "number_of v", standard]
-declare zpower_number_of_odd_number_of [simp]
-
-
-
-
-ML
-{*
-val numerals = thms"numerals";
-val numeral_ss = simpset() addsimps numerals;
-
-val nat_bin_arith_setup =
- Fast_Arith.map_data
-   (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
-     {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
-      inj_thms = inj_thms,
-      lessD = lessD, neqE = neqE,
-      simpset = simpset addsimps [Suc_nat_number_of, int_nat_number_of,
-                                  not_neg_number_of_Pls,
-                                  neg_number_of_Min,neg_number_of_BIT]})
-*}
-
-setup nat_bin_arith_setup
-
-(* Enable arith to deal with div/mod k where k is a numeral: *)
-declare split_div[of _ _ "number_of k", standard, arith_split]
-declare split_mod[of _ _ "number_of k", standard, arith_split]
-
-lemma nat_number_of_Pls: "Numeral0 = (0::nat)"
-  by (simp add: number_of_Pls nat_number_of_def)
-
-lemma nat_number_of_Min: "number_of Numeral.Min = (0::nat)"
-  apply (simp only: number_of_Min nat_number_of_def nat_zminus_int)
-  done
-
-lemma nat_number_of_BIT_1:
-  "number_of (w BIT bit.B1) =
-    (if neg (number_of w :: int) then 0
-     else let n = number_of w in Suc (n + n))"
-  apply (simp only: nat_number_of_def Let_def split: split_if)
-  apply (intro conjI impI)
-   apply (simp add: neg_nat neg_number_of_BIT)
-  apply (rule int_int_eq [THEN iffD1])
-  apply (simp only: not_neg_nat neg_number_of_BIT int_Suc zadd_int [symmetric] simp_thms)
-  apply (simp only: number_of_BIT zadd_assoc split: bit.split)
-  apply simp
-  done
-
-lemma nat_number_of_BIT_0:
-    "number_of (w BIT bit.B0) = (let n::nat = number_of w in n + n)"
-  apply (simp only: nat_number_of_def Let_def)
-  apply (cases "neg (number_of w :: int)")
-   apply (simp add: neg_nat neg_number_of_BIT)
-  apply (rule int_int_eq [THEN iffD1])
-  apply (simp only: not_neg_nat neg_number_of_BIT int_Suc zadd_int [symmetric] simp_thms)
-  apply (simp only: number_of_BIT zadd_assoc)
-  apply simp
-  done
-
-lemmas nat_number =
-  nat_number_of_Pls nat_number_of_Min
-  nat_number_of_BIT_1 nat_number_of_BIT_0
-
-lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
-  by (simp add: Let_def)
-
-lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring,recpower})"
-by (simp add: power_mult); 
-
-lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring,recpower})"
-by (simp add: power_mult power_Suc); 
-
-
-subsection{*Literal arithmetic and @{term of_nat}*}
-
-lemma of_nat_double:
-     "0 \<le> x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)"
-by (simp only: mult_2 nat_add_distrib of_nat_add) 
-
-lemma nat_numeral_m1_eq_0: "-1 = (0::nat)"
-by (simp only: nat_number_of_def)
-
-lemma of_nat_number_of_lemma:
-     "of_nat (number_of v :: nat) =  
-         (if 0 \<le> (number_of v :: int) 
-          then (number_of v :: 'a :: number_ring)
-          else 0)"
-by (simp add: int_number_of_def nat_number_of_def number_of_eq of_nat_nat);
-
-lemma of_nat_number_of_eq [simp]:
-     "of_nat (number_of v :: nat) =  
-         (if neg (number_of v :: int) then 0  
-          else (number_of v :: 'a :: number_ring))"
-by (simp only: of_nat_number_of_lemma neg_def, simp) 
-
-
-subsection {*Lemmas for the Combination and Cancellation Simprocs*}
-
-lemma nat_number_of_add_left:
-     "number_of v + (number_of v' + (k::nat)) =  
-         (if neg (number_of v :: int) then number_of v' + k  
-          else if neg (number_of v' :: int) then number_of v + k  
-          else number_of (v + v') + k)"
-by simp
-
-lemma nat_number_of_mult_left:
-     "number_of v * (number_of v' * (k::nat)) =  
-         (if neg (number_of v :: int) then 0
-          else number_of (v * v') * k)"
-by simp
-
-
-subsubsection{*For @{text combine_numerals}*}
-
-lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)"
-by (simp add: add_mult_distrib)
-
-
-subsubsection{*For @{text cancel_numerals}*}
-
-lemma nat_diff_add_eq1:
-     "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)"
-by (simp split add: nat_diff_split add: add_mult_distrib)
-
-lemma nat_diff_add_eq2:
-     "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))"
-by (simp split add: nat_diff_split add: add_mult_distrib)
-
-lemma nat_eq_add_iff1:
-     "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)"
-by (auto split add: nat_diff_split simp add: add_mult_distrib)
-
-lemma nat_eq_add_iff2:
-     "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)"
-by (auto split add: nat_diff_split simp add: add_mult_distrib)
-
-lemma nat_less_add_iff1:
-     "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)"
-by (auto split add: nat_diff_split simp add: add_mult_distrib)
-
-lemma nat_less_add_iff2:
-     "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)"
-by (auto split add: nat_diff_split simp add: add_mult_distrib)
-
-lemma nat_le_add_iff1:
-     "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)"
-by (auto split add: nat_diff_split simp add: add_mult_distrib)
-
-lemma nat_le_add_iff2:
-     "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)"
-by (auto split add: nat_diff_split simp add: add_mult_distrib)
-
-
-subsubsection{*For @{text cancel_numeral_factors} *}
-
-lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)"
-by auto
-
-lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m<n)"
-by auto
-
-lemma nat_mult_eq_cancel1: "(0::nat) < k ==> (k*m = k*n) = (m=n)"
-by auto
-
-lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)"
-by auto
-
-
-subsubsection{*For @{text cancel_factor} *}
-
-lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)"
-by auto
-
-lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m<n)"
-by auto
-
-lemma nat_mult_eq_cancel_disj: "(k*m = k*n) = (k = (0::nat) | m=n)"
-by auto
-
-lemma nat_mult_div_cancel_disj:
-     "(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
-by (simp add: nat_mult_div_cancel1)
-
-
-subsection {* legacy ML bindings *}
-
-ML
-{*
-val eq_nat_nat_iff = thm"eq_nat_nat_iff";
-val eq_nat_number_of = thm"eq_nat_number_of";
-val less_nat_number_of = thm"less_nat_number_of";
-val power2_eq_square = thm "power2_eq_square";
-val zero_le_power2 = thm "zero_le_power2";
-val zero_less_power2 = thm "zero_less_power2";
-val zero_eq_power2 = thm "zero_eq_power2";
-val abs_power2 = thm "abs_power2";
-val power2_abs = thm "power2_abs";
-val power2_minus = thm "power2_minus";
-val power_minus1_even = thm "power_minus1_even";
-val power_minus_even = thm "power_minus_even";
-val odd_power_less_zero = thm "odd_power_less_zero";
-val odd_0_le_power_imp_0_le = thm "odd_0_le_power_imp_0_le";
-
-val Suc_pred' = thm"Suc_pred'";
-val expand_Suc = thm"expand_Suc";
-val Suc_eq_add_numeral_1 = thm"Suc_eq_add_numeral_1";
-val Suc_eq_add_numeral_1_left = thm"Suc_eq_add_numeral_1_left";
-val add_eq_if = thm"add_eq_if";
-val mult_eq_if = thm"mult_eq_if";
-val power_eq_if = thm"power_eq_if";
-val eq_number_of_0 = thm"eq_number_of_0";
-val eq_0_number_of = thm"eq_0_number_of";
-val less_0_number_of = thm"less_0_number_of";
-val neg_imp_number_of_eq_0 = thm"neg_imp_number_of_eq_0";
-val eq_number_of_Suc = thm"eq_number_of_Suc";
-val Suc_eq_number_of = thm"Suc_eq_number_of";
-val less_number_of_Suc = thm"less_number_of_Suc";
-val less_Suc_number_of = thm"less_Suc_number_of";
-val le_number_of_Suc = thm"le_number_of_Suc";
-val le_Suc_number_of = thm"le_Suc_number_of";
-val eq_number_of_BIT_BIT = thm"eq_number_of_BIT_BIT";
-val eq_number_of_BIT_Pls = thm"eq_number_of_BIT_Pls";
-val eq_number_of_BIT_Min = thm"eq_number_of_BIT_Min";
-val eq_number_of_Pls_Min = thm"eq_number_of_Pls_Min";
-val of_nat_number_of_eq = thm"of_nat_number_of_eq";
-val nat_power_eq = thm"nat_power_eq";
-val power_nat_number_of = thm"power_nat_number_of";
-val zpower_number_of_even = thm"zpower_number_of_even";
-val zpower_number_of_odd = thm"zpower_number_of_odd";
-val nat_number_of_Pls = thm"nat_number_of_Pls";
-val nat_number_of_Min = thm"nat_number_of_Min";
-val Let_Suc = thm"Let_Suc";
-
-val nat_number = thms"nat_number";
-
-val nat_number_of_add_left = thm"nat_number_of_add_left";
-val nat_number_of_mult_left = thm"nat_number_of_mult_left";
-val left_add_mult_distrib = thm"left_add_mult_distrib";
-val nat_diff_add_eq1 = thm"nat_diff_add_eq1";
-val nat_diff_add_eq2 = thm"nat_diff_add_eq2";
-val nat_eq_add_iff1 = thm"nat_eq_add_iff1";
-val nat_eq_add_iff2 = thm"nat_eq_add_iff2";
-val nat_less_add_iff1 = thm"nat_less_add_iff1";
-val nat_less_add_iff2 = thm"nat_less_add_iff2";
-val nat_le_add_iff1 = thm"nat_le_add_iff1";
-val nat_le_add_iff2 = thm"nat_le_add_iff2";
-val nat_mult_le_cancel1 = thm"nat_mult_le_cancel1";
-val nat_mult_less_cancel1 = thm"nat_mult_less_cancel1";
-val nat_mult_eq_cancel1 = thm"nat_mult_eq_cancel1";
-val nat_mult_div_cancel1 = thm"nat_mult_div_cancel1";
-val nat_mult_le_cancel_disj = thm"nat_mult_le_cancel_disj";
-val nat_mult_less_cancel_disj = thm"nat_mult_less_cancel_disj";
-val nat_mult_eq_cancel_disj = thm"nat_mult_eq_cancel_disj";
-val nat_mult_div_cancel_disj = thm"nat_mult_div_cancel_disj";
-
-val power_minus_even = thm"power_minus_even";
-*}
-
-end
--- a/src/HOL/Integ/NatSimprocs.thy	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,391 +0,0 @@
-(*  Title:      HOL/NatSimprocs.thy
-    ID:         $Id$
-    Copyright   2003 TU Muenchen
-*)
-
-header {*Simprocs for the Naturals*}
-
-theory NatSimprocs
-imports NatBin
-uses "int_factor_simprocs.ML" "nat_simprocs.ML"
-begin
-
-setup nat_simprocs_setup
-
-subsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
-
-text{*Where K above is a literal*}
-
-lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
-by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
-
-text {*Now just instantiating @{text n} to @{text "number_of v"} does
-  the right simplification, but with some redundant inequality
-  tests.*}
-lemma neg_number_of_pred_iff_0:
-  "neg (number_of (Numeral.pred v)::int) = (number_of v = (0::nat))"
-apply (subgoal_tac "neg (number_of (Numeral.pred v)) = (number_of v < Suc 0) ")
-apply (simp only: less_Suc_eq_le le_0_eq)
-apply (subst less_number_of_Suc, simp)
-done
-
-text{*No longer required as a simprule because of the @{text inverse_fold}
-   simproc*}
-lemma Suc_diff_number_of:
-     "neg (number_of (uminus v)::int) ==>  
-      Suc m - (number_of v) = m - (number_of (Numeral.pred v))"
-apply (subst Suc_diff_eq_diff_pred)
-apply simp
-apply (simp del: nat_numeral_1_eq_1)
-apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] 
-                        neg_number_of_pred_iff_0)
-done
-
-lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
-by (simp add: numerals split add: nat_diff_split)
-
-
-subsection{*For @{term nat_case} and @{term nat_rec}*}
-
-lemma nat_case_number_of [simp]:
-     "nat_case a f (number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then a else f (nat pv))"
-by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
-
-lemma nat_case_add_eq_if [simp]:
-     "nat_case a f ((number_of v) + n) =  
-       (let pv = number_of (Numeral.pred v) in  
-         if neg pv then nat_case a f n else f (nat pv + n))"
-apply (subst add_eq_if)
-apply (simp split add: nat.split
-            del: nat_numeral_1_eq_1
-	    add: numeral_1_eq_Suc_0 [symmetric] Let_def 
-                 neg_imp_number_of_eq_0 neg_number_of_pred_iff_0)
-done
-
-lemma nat_rec_number_of [simp]:
-     "nat_rec a f (number_of v) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
-apply (case_tac " (number_of v) ::nat")
-apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
-apply (simp split add: split_if_asm)
-done
-
-lemma nat_rec_add_eq_if [simp]:
-     "nat_rec a f (number_of v + n) =  
-        (let pv = number_of (Numeral.pred v) in  
-         if neg pv then nat_rec a f n  
-                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
-apply (subst add_eq_if)
-apply (simp split add: nat.split
-            del: nat_numeral_1_eq_1
-            add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
-                 neg_number_of_pred_iff_0)
-done
-
-
-subsection{*Various Other Lemmas*}
-
-subsubsection{*Evens and Odds, for Mutilated Chess Board*}
-
-text{*Lemmas for specialist use, NOT as default simprules*}
-lemma nat_mult_2: "2 * z = (z+z::nat)"
-proof -
-  have "2*z = (1 + 1)*z" by simp
-  also have "... = z+z" by (simp add: left_distrib)
-  finally show ?thesis .
-qed
-
-lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
-by (subst mult_commute, rule nat_mult_2)
-
-text{*Case analysis on @{term "n<2"}*}
-lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
-by arith
-
-lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
-by arith
-
-lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
-by (simp add: nat_mult_2 [symmetric])
-
-lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
-apply (subgoal_tac "m mod 2 < 2")
-apply (erule less_2_cases [THEN disjE])
-apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
-done
-
-lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
-apply (subgoal_tac "m mod 2 < 2")
-apply (force simp del: mod_less_divisor, simp) 
-done
-
-subsubsection{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
-
-lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
-by simp
-
-lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
-by simp
-
-text{*Can be used to eliminate long strings of Sucs, but not by default*}
-lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
-by simp
-
-
-text{*These lemmas collapse some needless occurrences of Suc:
-    at least three Sucs, since two and fewer are rewritten back to Suc again!
-    We already have some rules to simplify operands smaller than 3.*}
-
-lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
-by (simp add: Suc3_eq_add_3)
-
-lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
-by (simp add: Suc3_eq_add_3)
-
-lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
-by (simp add: Suc3_eq_add_3)
-
-lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
-by (simp add: Suc3_eq_add_3)
-
-lemmas Suc_div_eq_add3_div_number_of =
-    Suc_div_eq_add3_div [of _ "number_of v", standard]
-declare Suc_div_eq_add3_div_number_of [simp]
-
-lemmas Suc_mod_eq_add3_mod_number_of =
-    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
-declare Suc_mod_eq_add3_mod_number_of [simp]
-
-
-
-subsection{*Special Simplification for Constants*}
-
-text{*These belong here, late in the development of HOL, to prevent their
-interfering with proofs of abstract properties of instances of the function
-@{term number_of}*}
-
-text{*These distributive laws move literals inside sums and differences.*}
-lemmas left_distrib_number_of = left_distrib [of _ _ "number_of v", standard]
-declare left_distrib_number_of [simp]
-
-lemmas right_distrib_number_of = right_distrib [of "number_of v", standard]
-declare right_distrib_number_of [simp]
-
-
-lemmas left_diff_distrib_number_of =
-    left_diff_distrib [of _ _ "number_of v", standard]
-declare left_diff_distrib_number_of [simp]
-
-lemmas right_diff_distrib_number_of =
-    right_diff_distrib [of "number_of v", standard]
-declare right_diff_distrib_number_of [simp]
-
-
-text{*These are actually for fields, like real: but where else to put them?*}
-lemmas zero_less_divide_iff_number_of =
-    zero_less_divide_iff [of "number_of w", standard]
-declare zero_less_divide_iff_number_of [simp]
-
-lemmas divide_less_0_iff_number_of =
-    divide_less_0_iff [of "number_of w", standard]
-declare divide_less_0_iff_number_of [simp]
-
-lemmas zero_le_divide_iff_number_of =
-    zero_le_divide_iff [of "number_of w", standard]
-declare zero_le_divide_iff_number_of [simp]
-
-lemmas divide_le_0_iff_number_of =
-    divide_le_0_iff [of "number_of w", standard]
-declare divide_le_0_iff_number_of [simp]
-
-
-(****
-IF times_divide_eq_right and times_divide_eq_left are removed as simprules,
-then these special-case declarations may be useful.
-
-text{*These simprules move numerals into numerators and denominators.*}
-lemma times_recip_eq_right [simp]: "a * (1/c) = a / (c::'a::field)"
-by (simp add: times_divide_eq)
-
-lemma times_recip_eq_left [simp]: "(1/c) * a = a / (c::'a::field)"
-by (simp add: times_divide_eq)
-
-lemmas times_divide_eq_right_number_of =
-    times_divide_eq_right [of "number_of w", standard]
-declare times_divide_eq_right_number_of [simp]
-
-lemmas times_divide_eq_right_number_of =
-    times_divide_eq_right [of _ _ "number_of w", standard]
-declare times_divide_eq_right_number_of [simp]
-
-lemmas times_divide_eq_left_number_of =
-    times_divide_eq_left [of _ "number_of w", standard]
-declare times_divide_eq_left_number_of [simp]
-
-lemmas times_divide_eq_left_number_of =
-    times_divide_eq_left [of _ _ "number_of w", standard]
-declare times_divide_eq_left_number_of [simp]
-
-****)
-
-text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
-  strange, but then other simprocs simplify the quotient.*}
-
-lemmas inverse_eq_divide_number_of =
-    inverse_eq_divide [of "number_of w", standard]
-declare inverse_eq_divide_number_of [simp]
-
-
-subsubsection{*These laws simplify inequalities, moving unary minus from a term
-into the literal.*}
-lemmas less_minus_iff_number_of =
-    less_minus_iff [of "number_of v", standard]
-declare less_minus_iff_number_of [simp]
-
-lemmas le_minus_iff_number_of =
-    le_minus_iff [of "number_of v", standard]
-declare le_minus_iff_number_of [simp]
-
-lemmas equation_minus_iff_number_of =
-    equation_minus_iff [of "number_of v", standard]
-declare equation_minus_iff_number_of [simp]
-
-
-lemmas minus_less_iff_number_of =
-    minus_less_iff [of _ "number_of v", standard]
-declare minus_less_iff_number_of [simp]
-
-lemmas minus_le_iff_number_of =
-    minus_le_iff [of _ "number_of v", standard]
-declare minus_le_iff_number_of [simp]
-
-lemmas minus_equation_iff_number_of =
-    minus_equation_iff [of _ "number_of v", standard]
-declare minus_equation_iff_number_of [simp]
-
-
-subsubsection{*To Simplify Inequalities Where One Side is the Constant 1*}
-
-lemma less_minus_iff_1 [simp]: 
-  fixes b::"'b::{ordered_idom,number_ring}" 
-  shows "(1 < - b) = (b < -1)"
-by auto
-
-lemma le_minus_iff_1 [simp]: 
-  fixes b::"'b::{ordered_idom,number_ring}" 
-  shows "(1 \<le> - b) = (b \<le> -1)"
-by auto
-
-lemma equation_minus_iff_1 [simp]: 
-  fixes b::"'b::number_ring" 
-  shows "(1 = - b) = (b = -1)"
-by (subst equation_minus_iff, auto) 
-
-lemma minus_less_iff_1 [simp]: 
-  fixes a::"'b::{ordered_idom,number_ring}" 
-  shows "(- a < 1) = (-1 < a)"
-by auto
-
-lemma minus_le_iff_1 [simp]: 
-  fixes a::"'b::{ordered_idom,number_ring}" 
-  shows "(- a \<le> 1) = (-1 \<le> a)"
-by auto
-
-lemma minus_equation_iff_1 [simp]: 
-  fixes a::"'b::number_ring" 
-  shows "(- a = 1) = (a = -1)"
-by (subst minus_equation_iff, auto) 
-
-
-subsubsection {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
-
-lemmas mult_less_cancel_left_number_of =
-    mult_less_cancel_left [of "number_of v", standard]
-declare mult_less_cancel_left_number_of [simp]
-
-lemmas mult_less_cancel_right_number_of =
-    mult_less_cancel_right [of _ "number_of v", standard]
-declare mult_less_cancel_right_number_of [simp]
-
-lemmas mult_le_cancel_left_number_of =
-    mult_le_cancel_left [of "number_of v", standard]
-declare mult_le_cancel_left_number_of [simp]
-
-lemmas mult_le_cancel_right_number_of =
-    mult_le_cancel_right [of _ "number_of v", standard]
-declare mult_le_cancel_right_number_of [simp]
-
-
-subsubsection {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
-
-lemmas le_divide_eq_number_of = le_divide_eq [of _ _ "number_of w", standard]
-declare le_divide_eq_number_of [simp]
-
-lemmas divide_le_eq_number_of = divide_le_eq [of _ "number_of w", standard]
-declare divide_le_eq_number_of [simp]
-
-lemmas less_divide_eq_number_of = less_divide_eq [of _ _ "number_of w", standard]
-declare less_divide_eq_number_of [simp]
-
-lemmas divide_less_eq_number_of = divide_less_eq [of _ "number_of w", standard]
-declare divide_less_eq_number_of [simp]
-
-lemmas eq_divide_eq_number_of = eq_divide_eq [of _ _ "number_of w", standard]
-declare eq_divide_eq_number_of [simp]
-
-lemmas divide_eq_eq_number_of = divide_eq_eq [of _ "number_of w", standard]
-declare divide_eq_eq_number_of [simp]
-
-
-
-subsection{*Optional Simplification Rules Involving Constants*}
-
-text{*Simplify quotients that are compared with a literal constant.*}
-
-lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard]
-lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard]
-lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard]
-lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard]
-lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard]
-lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard]
-
-
-text{*Not good as automatic simprules because they cause case splits.*}
-lemmas divide_const_simps =
-  le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of
-  divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of
-  le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
-
-subsubsection{*Division By @{text "-1"}*}
-
-lemma divide_minus1 [simp]:
-     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})" 
-by simp
-
-lemma minus1_divide [simp]:
-     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
-by (simp add: divide_inverse inverse_minus_eq)
-
-lemma half_gt_zero_iff:
-     "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
-by auto
-
-lemmas half_gt_zero = half_gt_zero_iff [THEN iffD2, standard]
-declare half_gt_zero [simp]
-
-(* The following lemma should appear in Divides.thy, but there the proof
-   doesn't work. *)
-
-lemma nat_dvd_not_less:
-  "[| 0 < m; m < n |] ==> \<not> n dvd (m::nat)"
-  by (unfold dvd_def) auto
-
-ML {*
-val divide_minus1 = @{thm divide_minus1};
-val minus1_divide = @{thm minus1_divide};
-*}
-
-end
--- a/src/HOL/Integ/Numeral.thy	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,685 +0,0 @@
-(*  Title:      HOL/Integ/Numeral.thy
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1994  University of Cambridge
-*)
-
-header {* Arithmetic on Binary Integers *}
-
-theory Numeral
-imports IntDef
-uses ("../Tools/numeral_syntax.ML")
-begin
-
-subsection {* Binary representation *}
-
-text {*
-  This formalization defines binary arithmetic in terms of the integers
-  rather than using a datatype. This avoids multiple representations (leading
-  zeroes, etc.)  See @{text "ZF/Integ/twos-compl.ML"}, function @{text
-  int_of_binary}, for the numerical interpretation.
-
-  The representation expects that @{text "(m mod 2)"} is 0 or 1,
-  even if m is negative;
-  For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
-  @{text "-5 = (-3)*2 + 1"}.
-*}
-
-datatype bit = B0 | B1
-
-text{*
-  Type @{typ bit} avoids the use of type @{typ bool}, which would make
-  all of the rewrite rules higher-order.
-*}
-
-definition
-  Pls :: int where
-  [code func del]:"Pls = 0"
-
-definition
-  Min :: int where
-  [code func del]:"Min = - 1"
-
-definition
-  Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where
-  [code func del]: "k BIT b = (case b of B0 \<Rightarrow> 0 | B1 \<Rightarrow> 1) + k + k"
-
-class number = type + -- {* for numeric types: nat, int, real, \dots *}
-  fixes number_of :: "int \<Rightarrow> 'a"
-
-syntax
-  "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
-
-use "../Tools/numeral_syntax.ML"
-setup NumeralSyntax.setup
-
-abbreviation
-  "Numeral0 \<equiv> number_of Pls"
-
-abbreviation
-  "Numeral1 \<equiv> number_of (Pls BIT B1)"
-
-lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)"
-  -- {* Unfold all @{text let}s involving constants *}
-  unfolding Let_def ..
-
-lemma Let_0 [simp]: "Let 0 f = f 0"
-  unfolding Let_def ..
-
-lemma Let_1 [simp]: "Let 1 f = f 1"
-  unfolding Let_def ..
-
-definition
-  succ :: "int \<Rightarrow> int" where
-  [code func del]: "succ k = k + 1"
-
-definition
-  pred :: "int \<Rightarrow> int" where
-  [code func del]: "pred k = k - 1"
-
-lemmas
-  max_number_of [simp] = max_def
-    [of "number_of u" "number_of v", standard, simp]
-and
-  min_number_of [simp] = min_def 
-    [of "number_of u" "number_of v", standard, simp]
-  -- {* unfolding @{text minx} and @{text max} on numerals *}
-
-lemmas numeral_simps = 
-  succ_def pred_def Pls_def Min_def Bit_def
-
-text {* Removal of leading zeroes *}
-
-lemma Pls_0_eq [simp, normal post]:
-  "Pls BIT B0 = Pls"
-  unfolding numeral_simps by simp
-
-lemma Min_1_eq [simp, normal post]:
-  "Min BIT B1 = Min"
-  unfolding numeral_simps by simp
-
-
-subsection {* The Functions @{term succ}, @{term pred} and @{term uminus} *}
-
-lemma succ_Pls [simp]:
-  "succ Pls = Pls BIT B1"
-  unfolding numeral_simps by simp
-
-lemma succ_Min [simp]:
-  "succ Min = Pls"
-  unfolding numeral_simps by simp
-
-lemma succ_1 [simp]:
-  "succ (k BIT B1) = succ k BIT B0"
-  unfolding numeral_simps by simp
-
-lemma succ_0 [simp]:
-  "succ (k BIT B0) = k BIT B1"
-  unfolding numeral_simps by simp
-
-lemma pred_Pls [simp]:
-  "pred Pls = Min"
-  unfolding numeral_simps by simp
-
-lemma pred_Min [simp]:
-  "pred Min = Min BIT B0"
-  unfolding numeral_simps by simp
-
-lemma pred_1 [simp]:
-  "pred (k BIT B1) = k BIT B0"
-  unfolding numeral_simps by simp
-
-lemma pred_0 [simp]:
-  "pred (k BIT B0) = pred k BIT B1"
-  unfolding numeral_simps by simp 
-
-lemma minus_Pls [simp]:
-  "- Pls = Pls"
-  unfolding numeral_simps by simp 
-
-lemma minus_Min [simp]:
-  "- Min = Pls BIT B1"
-  unfolding numeral_simps by simp 
-
-lemma minus_1 [simp]:
-  "- (k BIT B1) = pred (- k) BIT B1"
-  unfolding numeral_simps by simp 
-
-lemma minus_0 [simp]:
-  "- (k BIT B0) = (- k) BIT B0"
-  unfolding numeral_simps by simp 
-
-
-subsection {*
-  Binary Addition and Multiplication: @{term "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"}
-    and @{term "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"}
-*}
-
-lemma add_Pls [simp]:
-  "Pls + k = k"
-  unfolding numeral_simps by simp 
-
-lemma add_Min [simp]:
-  "Min + k = pred k"
-  unfolding numeral_simps by simp
-
-lemma add_BIT_11 [simp]:
-  "(k BIT B1) + (l BIT B1) = (k + succ l) BIT B0"
-  unfolding numeral_simps by simp
-
-lemma add_BIT_10 [simp]:
-  "(k BIT B1) + (l BIT B0) = (k + l) BIT B1"
-  unfolding numeral_simps by simp
-
-lemma add_BIT_0 [simp]:
-  "(k BIT B0) + (l BIT b) = (k + l) BIT b"
-  unfolding numeral_simps by simp 
-
-lemma add_Pls_right [simp]:
-  "k + Pls = k"
-  unfolding numeral_simps by simp 
-
-lemma add_Min_right [simp]:
-  "k + Min = pred k"
-  unfolding numeral_simps by simp 
-
-lemma mult_Pls [simp]:
-  "Pls * w = Pls"
-  unfolding numeral_simps by simp 
-
-lemma mult_Min [simp]:
-  "Min * k = - k"
-  unfolding numeral_simps by simp 
-
-lemma mult_num1 [simp]:
-  "(k BIT B1) * l = ((k * l) BIT B0) + l"
-  unfolding numeral_simps int_distrib by simp 
-
-lemma mult_num0 [simp]:
-  "(k BIT B0) * l = (k * l) BIT B0"
-  unfolding numeral_simps int_distrib by simp 
-
-
-
-subsection {* Converting Numerals to Rings: @{term number_of} *}
-
-axclass number_ring \<subseteq> number, comm_ring_1
-  number_of_eq: "number_of k = of_int k"
-
-text {* self-embedding of the intergers *}
-
-instance int :: number_ring
-  int_number_of_def: "number_of w \<equiv> of_int w"
-  by intro_classes (simp only: int_number_of_def)
-
-lemmas [code func del] = int_number_of_def
-
-lemma number_of_is_id:
-  "number_of (k::int) = k"
-  unfolding int_number_of_def by simp
-
-lemma number_of_succ:
-  "number_of (succ k) = (1 + number_of k ::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma number_of_pred:
-  "number_of (pred w) = (- 1 + number_of w ::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma number_of_minus:
-  "number_of (uminus w) = (- (number_of w)::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma number_of_add:
-  "number_of (v + w) = (number_of v + number_of w::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma number_of_mult:
-  "number_of (v * w) = (number_of v * number_of w::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-text {*
-  The correctness of shifting.
-  But it doesn't seem to give a measurable speed-up.
-*}
-
-lemma double_number_of_BIT:
-  "(1 + 1) * number_of w = (number_of (w BIT B0) ::'a::number_ring)"
-  unfolding number_of_eq numeral_simps left_distrib by simp
-
-text {*
-  Converting numerals 0 and 1 to their abstract versions.
-*}
-
-lemma numeral_0_eq_0 [simp]:
-  "Numeral0 = (0::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma numeral_1_eq_1 [simp]:
-  "Numeral1 = (1::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-text {*
-  Special-case simplification for small constants.
-*}
-
-text{*
-  Unary minus for the abstract constant 1. Cannot be inserted
-  as a simprule until later: it is @{text number_of_Min} re-oriented!
-*}
-
-lemma numeral_m1_eq_minus_1:
-  "(-1::'a::number_ring) = - 1"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma mult_minus1 [simp]:
-  "-1 * z = -(z::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma mult_minus1_right [simp]:
-  "z * -1 = -(z::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-(*Negation of a coefficient*)
-lemma minus_number_of_mult [simp]:
-   "- (number_of w) * z = number_of (uminus w) * (z::'a::number_ring)"
-   unfolding number_of_eq by simp
-
-text {* Subtraction *}
-
-lemma diff_number_of_eq:
-  "number_of v - number_of w =
-    (number_of (v + uminus w)::'a::number_ring)"
-  unfolding number_of_eq by simp
-
-lemma number_of_Pls:
-  "number_of Pls = (0::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma number_of_Min:
-  "number_of Min = (- 1::'a::number_ring)"
-  unfolding number_of_eq numeral_simps by simp
-
-lemma number_of_BIT:
-  "number_of(w BIT x) = (case x of B0 => 0 | B1 => (1::'a::number_ring))
-    + (number_of w) + (number_of w)"
-  unfolding number_of_eq numeral_simps by (simp split: bit.split)
-
-
-subsection {* Equality of Binary Numbers *}
-
-text {* First version by Norbert Voelker *}
-
-lemma eq_number_of_eq:
-  "((number_of x::'a::number_ring) = number_of y) =
-   iszero (number_of (x + uminus y) :: 'a)"
-  unfolding iszero_def number_of_add number_of_minus
-  by (simp add: compare_rls)
-
-lemma iszero_number_of_Pls:
-  "iszero ((number_of Pls)::'a::number_ring)"
-  unfolding iszero_def numeral_0_eq_0 ..
-
-lemma nonzero_number_of_Min:
-  "~ iszero ((number_of Min)::'a::number_ring)"
-  unfolding iszero_def numeral_m1_eq_minus_1 by simp
-
-
-subsection {* Comparisons, for Ordered Rings *}
-
-lemma double_eq_0_iff:
-  "(a + a = 0) = (a = (0::'a::ordered_idom))"
-proof -
-  have "a + a = (1 + 1) * a" unfolding left_distrib by simp
-  with zero_less_two [where 'a = 'a]
-  show ?thesis by force
-qed
-
-lemma le_imp_0_less: 
-  assumes le: "0 \<le> z"
-  shows "(0::int) < 1 + z"
-proof -
-  have "0 \<le> z" .
-  also have "... < z + 1" by (rule less_add_one) 
-  also have "... = 1 + z" by (simp add: add_ac)
-  finally show "0 < 1 + z" .
-qed
-
-lemma odd_nonzero:
-  "1 + z + z \<noteq> (0::int)";
-proof (cases z rule: int_cases)
-  case (nonneg n)
-  have le: "0 \<le> z+z" by (simp add: nonneg add_increasing) 
-  thus ?thesis using  le_imp_0_less [OF le]
-    by (auto simp add: add_assoc) 
-next
-  case (neg n)
-  show ?thesis
-  proof
-    assume eq: "1 + z + z = 0"
-    have "0 < 1 + (int n + int n)"
-      by (simp add: le_imp_0_less add_increasing) 
-    also have "... = - (1 + z + z)" 
-      by (simp add: neg add_assoc [symmetric]) 
-    also have "... = 0" by (simp add: eq) 
-    finally have "0<0" ..
-    thus False by blast
-  qed
-qed
-
-text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
-
-lemma Ints_double_eq_0_iff:
-  assumes in_Ints: "a \<in> Ints"
-  shows "(a + a = 0) = (a = (0::'a::ring_char_0))"
-proof -
-  from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] .
-  then obtain z where a: "a = of_int z" ..
-  show ?thesis
-  proof
-    assume "a = 0"
-    thus "a + a = 0" by simp
-  next
-    assume eq: "a + a = 0"
-    hence "of_int (z + z) = (of_int 0 :: 'a)" by (simp add: a)
-    hence "z + z = 0" by (simp only: of_int_eq_iff)
-    hence "z = 0" by (simp only: double_eq_0_iff)
-    thus "a = 0" by (simp add: a)
-  qed
-qed
-
-lemma Ints_odd_nonzero:
-  assumes in_Ints: "a \<in> Ints"
-  shows "1 + a + a \<noteq> (0::'a::ring_char_0)"
-proof -
-  from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] .
-  then obtain z where a: "a = of_int z" ..
-  show ?thesis
-  proof
-    assume eq: "1 + a + a = 0"
-    hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a)
-    hence "1 + z + z = 0" by (simp only: of_int_eq_iff)
-    with odd_nonzero show False by blast
-  qed
-qed 
-
-lemma Ints_number_of:
-  "(number_of w :: 'a::number_ring) \<in> Ints"
-  unfolding number_of_eq Ints_def by simp
-
-lemma iszero_number_of_BIT:
-  "iszero (number_of (w BIT x)::'a) = 
-   (x = B0 \<and> iszero (number_of w::'a::{ring_char_0,number_ring}))"
-  by (simp add: iszero_def number_of_eq numeral_simps Ints_double_eq_0_iff 
-    Ints_odd_nonzero Ints_def split: bit.split)
-
-lemma iszero_number_of_0:
-  "iszero (number_of (w BIT B0) :: 'a::{ring_char_0,number_ring}) = 
-  iszero (number_of w :: 'a)"
-  by (simp only: iszero_number_of_BIT simp_thms)
-
-lemma iszero_number_of_1:
-  "~ iszero (number_of (w BIT B1)::'a::{ring_char_0,number_ring})"
-  by (simp add: iszero_number_of_BIT) 
-
-
-subsection {* The Less-Than Relation *}
-
-lemma less_number_of_eq_neg:
-  "((number_of x::'a::{ordered_idom,number_ring}) < number_of y)
-  = neg (number_of (x + uminus y) :: 'a)"
-apply (subst less_iff_diff_less_0) 
-apply (simp add: neg_def diff_minus number_of_add number_of_minus)
-done
-
-text {*
-  If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
-  @{term Numeral0} IS @{term "number_of Pls"}
-*}
-
-lemma not_neg_number_of_Pls:
-  "~ neg (number_of Pls ::'a::{ordered_idom,number_ring})"
-  by (simp add: neg_def numeral_0_eq_0)
-
-lemma neg_number_of_Min:
-  "neg (number_of Min ::'a::{ordered_idom,number_ring})"
-  by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1)
-
-lemma double_less_0_iff:
-  "(a + a < 0) = (a < (0::'a::ordered_idom))"
-proof -
-  have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib)
-  also have "... = (a < 0)"
-    by (simp add: mult_less_0_iff zero_less_two 
-                  order_less_not_sym [OF zero_less_two]) 
-  finally show ?thesis .
-qed
-
-lemma odd_less_0:
-  "(1 + z + z < 0) = (z < (0::int))";
-proof (cases z rule: int_cases)
-  case (nonneg n)
-  thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing
-                             le_imp_0_less [THEN order_less_imp_le])  
-next
-  case (neg n)
-  thus ?thesis by (simp del: int_Suc
-    add: int_Suc0_eq_1 [symmetric] zadd_int compare_rls)
-qed
-
-text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
-
-lemma Ints_odd_less_0: 
-  assumes in_Ints: "a \<in> Ints"
-  shows "(1 + a + a < 0) = (a < (0::'a::ordered_idom))";
-proof -
-  from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] .
-  then obtain z where a: "a = of_int z" ..
-  hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))"
-    by (simp add: a)
-  also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0)
-  also have "... = (a < 0)" by (simp add: a)
-  finally show ?thesis .
-qed
-
-lemma neg_number_of_BIT:
-  "neg (number_of (w BIT x)::'a) = 
-  neg (number_of w :: 'a::{ordered_idom,number_ring})"
-  by (simp add: neg_def number_of_eq numeral_simps double_less_0_iff
-    Ints_odd_less_0 Ints_def split: bit.split)
-
-
-text {* Less-Than or Equals *}
-
-text {* Reduces @{term "a\<le>b"} to @{term "~ (b<a)"} for ALL numerals. *}
-
-lemmas le_number_of_eq_not_less =
-  linorder_not_less [of "number_of w" "number_of v", symmetric, 
-  standard]
-
-lemma le_number_of_eq:
-    "((number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y)
-     = (~ (neg (number_of (y + uminus x) :: 'a)))"
-by (simp add: le_number_of_eq_not_less less_number_of_eq_neg)
-
-
-text {* Absolute value (@{term abs}) *}
-
-lemma abs_number_of:
-  "abs(number_of x::'a::{ordered_idom,number_ring}) =
-   (if number_of x < (0::'a) then -number_of x else number_of x)"
-  by (simp add: abs_if)
-
-
-text {* Re-orientation of the equation nnn=x *}
-
-lemma number_of_reorient:
-  "(number_of w = x) = (x = number_of w)"
-  by auto
-
-
-subsection {* Simplification of arithmetic operations on integer constants. *}
-
-lemmas arith_extra_simps [standard, simp] =
-  number_of_add [symmetric]
-  number_of_minus [symmetric] numeral_m1_eq_minus_1 [symmetric]
-  number_of_mult [symmetric]
-  diff_number_of_eq abs_number_of 
-
-text {*
-  For making a minimal simpset, one must include these default simprules.
-  Also include @{text simp_thms}.
-*}
-
-lemmas arith_simps = 
-  bit.distinct
-  Pls_0_eq Min_1_eq
-  pred_Pls pred_Min pred_1 pred_0
-  succ_Pls succ_Min succ_1 succ_0
-  add_Pls add_Min add_BIT_0 add_BIT_10 add_BIT_11
-  minus_Pls minus_Min minus_1 minus_0
-  mult_Pls mult_Min mult_num1 mult_num0 
-  add_Pls_right add_Min_right
-  abs_zero abs_one arith_extra_simps
-
-text {* Simplification of relational operations *}
-
-lemmas rel_simps [simp] = 
-  eq_number_of_eq iszero_number_of_Pls nonzero_number_of_Min
-  iszero_number_of_0 iszero_number_of_1
-  less_number_of_eq_neg
-  not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1
-  neg_number_of_Min neg_number_of_BIT
-  le_number_of_eq
-
-
-subsection {* Simplification of arithmetic when nested to the right. *}
-
-lemma add_number_of_left [simp]:
-  "number_of v + (number_of w + z) =
-   (number_of(v + w) + z::'a::number_ring)"
-  by (simp add: add_assoc [symmetric])
-
-lemma mult_number_of_left [simp]:
-  "number_of v * (number_of w * z) =
-   (number_of(v * w) * z::'a::number_ring)"
-  by (simp add: mult_assoc [symmetric])
-
-lemma add_number_of_diff1:
-  "number_of v + (number_of w - c) = 
-  number_of(v + w) - (c::'a::number_ring)"
-  by (simp add: diff_minus add_number_of_left)
-
-lemma add_number_of_diff2 [simp]:
-  "number_of v + (c - number_of w) =
-   number_of (v + uminus w) + (c::'a::number_ring)"
-apply (subst diff_number_of_eq [symmetric])
-apply (simp only: compare_rls)
-done
-
-
-subsection {* Configuration of the code generator *}
-
-instance int :: eq ..
-
-code_datatype Pls Min Bit "number_of \<Colon> int \<Rightarrow> int"
-
-definition
-  int_aux :: "int \<Rightarrow> nat \<Rightarrow> int" where
-  "int_aux i n = (i + int n)"
-
-lemma [code]:
-  "int_aux i 0 = i"
-  "int_aux i (Suc n) = int_aux (i + 1) n" -- {* tail recursive *}
-  by (simp add: int_aux_def)+
-
-lemma [code]:
-  "int n = int_aux 0 n"
-  by (simp add: int_aux_def)
-
-definition
-  nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat" where
-  "nat_aux n i = (n + nat i)"
-
-lemma [code]: "nat_aux n i = (if i <= 0 then n else nat_aux (Suc n) (i - 1))"
-  -- {* tail recursive *}
-  by (auto simp add: nat_aux_def nat_eq_iff linorder_not_le order_less_imp_le
-    dest: zless_imp_add1_zle)
-
-lemma [code]: "nat i = nat_aux 0 i"
-  by (simp add: nat_aux_def)
-
-lemma zero_is_num_zero [code func, code inline, symmetric, normal post]:
-  "(0\<Colon>int) = number_of Numeral.Pls" 
-  by simp
-
-lemma one_is_num_one [code func, code inline, symmetric, normal post]:
-  "(1\<Colon>int) = number_of (Numeral.Pls BIT bit.B1)" 
-  by simp 
-
-code_modulename SML
-  IntDef Integer
-
-code_modulename OCaml
-  IntDef Integer
-
-code_modulename Haskell
-  IntDef Integer
-
-code_modulename SML
-  Numeral Integer
-
-code_modulename OCaml
-  Numeral Integer
-
-code_modulename Haskell
-  Numeral Integer
-
-(*FIXME: the IntInf.fromInt below hides a dependence on fixed-precision ints!*)
-
-types_code
-  "int" ("int")
-attach (term_of) {*
-val term_of_int = HOLogic.mk_number HOLogic.intT o IntInf.fromInt;
-*}
-attach (test) {*
-fun gen_int i = one_of [~1, 1] * random_range 0 i;
-*}
-
-setup {*
-let
-
-fun number_of_codegen thy defs gr dep module b (Const (@{const_name Numeral.number_of}, Type ("fun", [_, T])) $ t) =
-      if T = HOLogic.intT then
-        (SOME (fst (Codegen.invoke_tycodegen thy defs dep module false (gr, T)),
-          (Pretty.str o IntInf.toString o HOLogic.dest_numeral) t) handle TERM _ => NONE)
-      else if T = HOLogic.natT then
-        SOME (Codegen.invoke_codegen thy defs dep module b (gr,
-          Const ("IntDef.nat", HOLogic.intT --> HOLogic.natT) $
-            (Const (@{const_name Numeral.number_of}, HOLogic.intT --> HOLogic.intT) $ t)))
-      else NONE
-  | number_of_codegen _ _ _ _ _ _ _ = NONE;
-
-in
-
-Codegen.add_codegen "number_of_codegen" number_of_codegen
-
-end
-*}
-
-consts_code
-  "0 :: int"                   ("0")
-  "1 :: int"                   ("1")
-  "uminus :: int => int"       ("~")
-  "op + :: int => int => int"  ("(_ +/ _)")
-  "op * :: int => int => int"  ("(_ */ _)")
-  "op \<le> :: int => int => bool" ("(_ <=/ _)")
-  "op < :: int => int => bool" ("(_ </ _)")
-
-quickcheck_params [default_type = int]
-
-(*setup continues in theory Presburger*)
-
-hide (open) const Pls Min B0 B1 succ pred
-
-end
--- a/src/HOL/Integ/int_arith1.ML	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,610 +0,0 @@
-(*  Title:      HOL/Integ/int_arith1.ML
-    ID:         $Id$
-    Authors:    Larry Paulson and Tobias Nipkow
-
-Simprocs and decision procedure for linear arithmetic.
-*)
-
-(** Misc ML bindings **)
-
-val succ_Pls = thm "succ_Pls";
-val succ_Min = thm "succ_Min";
-val succ_1 = thm "succ_1";
-val succ_0 = thm "succ_0";
-
-val pred_Pls = thm "pred_Pls";
-val pred_Min = thm "pred_Min";
-val pred_1 = thm "pred_1";
-val pred_0 = thm "pred_0";
-
-val minus_Pls = thm "minus_Pls";
-val minus_Min = thm "minus_Min";
-val minus_1 = thm "minus_1";
-val minus_0 = thm "minus_0";
-
-val add_Pls = thm "add_Pls";
-val add_Min = thm "add_Min";
-val add_BIT_11 = thm "add_BIT_11";
-val add_BIT_10 = thm "add_BIT_10";
-val add_BIT_0 = thm "add_BIT_0";
-val add_Pls_right = thm "add_Pls_right";
-val add_Min_right = thm "add_Min_right";
-
-val mult_Pls = thm "mult_Pls";
-val mult_Min = thm "mult_Min";
-val mult_num1 = thm "mult_num1";
-val mult_num0 = thm "mult_num0";
-
-val neg_def = thm "neg_def";
-val iszero_def = thm "iszero_def";
-
-val number_of_succ = thm "number_of_succ";
-val number_of_pred = thm "number_of_pred";
-val number_of_minus = thm "number_of_minus";
-val number_of_add = thm "number_of_add";
-val diff_number_of_eq = thm "diff_number_of_eq";
-val number_of_mult = thm "number_of_mult";
-val double_number_of_BIT = thm "double_number_of_BIT";
-val numeral_0_eq_0 = thm "numeral_0_eq_0";
-val numeral_1_eq_1 = thm "numeral_1_eq_1";
-val numeral_m1_eq_minus_1 = thm "numeral_m1_eq_minus_1";
-val mult_minus1 = thm "mult_minus1";
-val mult_minus1_right = thm "mult_minus1_right";
-val minus_number_of_mult = thm "minus_number_of_mult";
-val zero_less_nat_eq = thm "zero_less_nat_eq";
-val eq_number_of_eq = thm "eq_number_of_eq";
-val iszero_number_of_Pls = thm "iszero_number_of_Pls";
-val nonzero_number_of_Min = thm "nonzero_number_of_Min";
-val iszero_number_of_BIT = thm "iszero_number_of_BIT";
-val iszero_number_of_0 = thm "iszero_number_of_0";
-val iszero_number_of_1 = thm "iszero_number_of_1";
-val less_number_of_eq_neg = thm "less_number_of_eq_neg";
-val le_number_of_eq = thm "le_number_of_eq";
-val not_neg_number_of_Pls = thm "not_neg_number_of_Pls";
-val neg_number_of_Min = thm "neg_number_of_Min";
-val neg_number_of_BIT = thm "neg_number_of_BIT";
-val le_number_of_eq_not_less = thm "le_number_of_eq_not_less";
-val abs_number_of = thm "abs_number_of";
-val number_of_reorient = thm "number_of_reorient";
-val add_number_of_left = thm "add_number_of_left";
-val mult_number_of_left = thm "mult_number_of_left";
-val add_number_of_diff1 = thm "add_number_of_diff1";
-val add_number_of_diff2 = thm "add_number_of_diff2";
-val less_iff_diff_less_0 = thm "less_iff_diff_less_0";
-val eq_iff_diff_eq_0 = thm "eq_iff_diff_eq_0";
-val le_iff_diff_le_0 = thm "le_iff_diff_le_0";
-
-val arith_extra_simps = thms "arith_extra_simps";
-val arith_simps = thms "arith_simps";
-val rel_simps = thms "rel_simps";
-
-val zless_imp_add1_zle = thm "zless_imp_add1_zle";
-
-val combine_common_factor = thm "combine_common_factor";
-val eq_add_iff1 = thm "eq_add_iff1";
-val eq_add_iff2 = thm "eq_add_iff2";
-val less_add_iff1 = thm "less_add_iff1";
-val less_add_iff2 = thm "less_add_iff2";
-val le_add_iff1 = thm "le_add_iff1";
-val le_add_iff2 = thm "le_add_iff2";
-
-val arith_special = thms "arith_special";
-
-structure Int_Numeral_Base_Simprocs =
-  struct
-  fun prove_conv tacs ctxt (_: thm list) (t, u) =
-    if t aconv u then NONE
-    else
-      let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))
-      in SOME (Goal.prove ctxt [] [] eq (K (EVERY tacs))) end
-
-  fun prove_conv_nohyps tacs sg = prove_conv tacs sg [];
-
-  fun prep_simproc (name, pats, proc) =
-    Simplifier.simproc (the_context()) name pats proc;
-
-  fun is_numeral (Const(@{const_name Numeral.number_of}, _) $ w) = true
-    | is_numeral _ = false
-
-  fun simplify_meta_eq f_number_of_eq f_eq =
-      mk_meta_eq ([f_eq, f_number_of_eq] MRS trans)
-
-  (*reorientation simprules using ==, for the following simproc*)
-  val meta_zero_reorient = zero_reorient RS eq_reflection
-  val meta_one_reorient = one_reorient RS eq_reflection
-  val meta_number_of_reorient = number_of_reorient RS eq_reflection
-
-  (*reorientation simplification procedure: reorients (polymorphic) 
-    0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
-  fun reorient_proc sg _ (_ $ t $ u) =
-    case u of
-	Const(@{const_name HOL.zero}, _) => NONE
-      | Const(@{const_name HOL.one}, _) => NONE
-      | Const(@{const_name Numeral.number_of}, _) $ _ => NONE
-      | _ => SOME (case t of
-		  Const(@{const_name HOL.zero}, _) => meta_zero_reorient
-		| Const(@{const_name HOL.one}, _) => meta_one_reorient
-		| Const(@{const_name Numeral.number_of}, _) $ _ => meta_number_of_reorient)
-
-  val reorient_simproc = 
-      prep_simproc ("reorient_simproc", ["0=x", "1=x", "number_of w = x"], reorient_proc)
-
-  end;
-
-
-Addsimprocs [Int_Numeral_Base_Simprocs.reorient_simproc];
-
-
-structure Int_Numeral_Simprocs =
-struct
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in Int_Numeral_Base_Simprocs
-  isn't complicated by the abstract 0 and 1.*)
-val numeral_syms = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym];
-
-(** New term ordering so that AC-rewriting brings numerals to the front **)
-
-(*Order integers by absolute value and then by sign. The standard integer
-  ordering is not well-founded.*)
-fun num_ord (i,j) =
-      (case IntInf.compare (IntInf.abs i, IntInf.abs j) of
-            EQUAL => int_ord (IntInf.sign i, IntInf.sign j) 
-          | ord => ord);
-
-(*This resembles Term.term_ord, but it puts binary numerals before other
-  non-atomic terms.*)
-local open Term 
-in 
-fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) =
-      (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
-  | numterm_ord
-     (Const(@{const_name Numeral.number_of}, _) $ v, Const(@{const_name Numeral.number_of}, _) $ w) =
-     num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
-  | numterm_ord (Const(@{const_name Numeral.number_of}, _) $ _, _) = LESS
-  | numterm_ord (_, Const(@{const_name Numeral.number_of}, _) $ _) = GREATER
-  | numterm_ord (t, u) =
-      (case int_ord (size_of_term t, size_of_term u) of
-        EQUAL =>
-          let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
-            (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
-          end
-      | ord => ord)
-and numterms_ord (ts, us) = list_ord numterm_ord (ts, us)
-end;
-
-fun numtermless tu = (numterm_ord tu = LESS);
-
-(*Defined in this file, but perhaps needed only for Int_Numeral_Base_Simprocs of type nat.*)
-val num_ss = HOL_ss settermless numtermless;
-
-
-(** Utilities **)
-
-fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n;
-
-fun find_first_numeral past (t::terms) =
-        ((snd (HOLogic.dest_number t), rev past @ terms)
-         handle TERM _ => find_first_numeral (t::past) terms)
-  | find_first_numeral past [] = raise TERM("find_first_numeral", []);
-
-val mk_plus = HOLogic.mk_binop @{const_name HOL.plus};
-
-fun mk_minus t = 
-  let val T = Term.fastype_of t
-  in Const (@{const_name HOL.uminus}, T --> T) $ t
-  end;
-
-(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
-fun mk_sum T []        = mk_number T 0
-  | mk_sum T [t,u]     = mk_plus (t, u)
-  | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
-
-(*this version ALWAYS includes a trailing zero*)
-fun long_mk_sum T []        = mk_number T 0
-  | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
-
-val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT;
-
-(*decompose additions AND subtractions as a sum*)
-fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) =
-        dest_summing (pos, t, dest_summing (pos, u, ts))
-  | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) =
-        dest_summing (pos, t, dest_summing (not pos, u, ts))
-  | dest_summing (pos, t, ts) =
-        if pos then t::ts else mk_minus t :: ts;
-
-fun dest_sum t = dest_summing (true, t, []);
-
-val mk_diff = HOLogic.mk_binop @{const_name HOL.minus};
-val dest_diff = HOLogic.dest_bin @{const_name HOL.minus} Term.dummyT;
-
-val mk_times = HOLogic.mk_binop @{const_name HOL.times};
-
-fun mk_prod T = 
-  let val one = mk_number T 1
-  fun mk [] = one
-    | mk [t] = t
-    | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts)
-  in mk end;
-
-(*This version ALWAYS includes a trailing one*)
-fun long_mk_prod T []        = mk_number T 1
-  | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts);
-
-val dest_times = HOLogic.dest_bin @{const_name HOL.times} Term.dummyT;
-
-fun dest_prod t =
-      let val (t,u) = dest_times t
-      in  dest_prod t @ dest_prod u  end
-      handle TERM _ => [t];
-
-(*DON'T do the obvious simplifications; that would create special cases*)
-fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t);
-
-(*Express t as a product of (possibly) a numeral with other sorted terms*)
-fun dest_coeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_coeff (~sign) t
-  | dest_coeff sign t =
-    let val ts = sort Term.term_ord (dest_prod t)
-        val (n, ts') = find_first_numeral [] ts
-                          handle TERM _ => (1, ts)
-    in (sign*n, mk_prod (Term.fastype_of t) ts') end;
-
-(*Find first coefficient-term THAT MATCHES u*)
-fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
-  | find_first_coeff past u (t::terms) =
-        let val (n,u') = dest_coeff 1 t
-        in  if u aconv u' then (n, rev past @ terms)
-                          else find_first_coeff (t::past) u terms
-        end
-        handle TERM _ => find_first_coeff (t::past) u terms;
-
-(*Fractions as pairs of ints. Can't use Rat.rat because the representation
-  needs to preserve negative values in the denominator.*)
-fun mk_frac (p, q : IntInf.int) = if q = 0 then raise Div else (p, q);
-
-(*Don't reduce fractions; sums must be proved by rule add_frac_eq.
-  Fractions are reduced later by the cancel_numeral_factor simproc.*)
-fun add_frac ((p1 : IntInf.int, q1 : IntInf.int), (p2, q2)) = (p1 * q2 + p2 * q1, q1 * q2);
-
-val mk_divide = HOLogic.mk_binop @{const_name HOL.divide};
-
-(*Build term (p / q) * t*)
-fun mk_fcoeff ((p, q), t) =
-  let val T = Term.fastype_of t
-  in  mk_times (mk_divide (mk_number T p, mk_number T q), t) end;
-
-(*Express t as a product of a fraction with other sorted terms*)
-fun dest_fcoeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_fcoeff (~sign) t
-  | dest_fcoeff sign (Const (@{const_name HOL.divide}, _) $ t $ u) =
-    let val (p, t') = dest_coeff sign t
-        val (q, u') = dest_coeff 1 u
-    in  (mk_frac (p, q), mk_divide (t', u')) end
-  | dest_fcoeff sign t =
-    let val (p, t') = dest_coeff sign t
-        val T = Term.fastype_of t
-    in  (mk_frac (p, 1), mk_divide (t', mk_number T 1)) end;
-
-
-(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
-val add_0s =  thms "add_0s";
-val mult_1s = thms "mult_1s";
-
-(*Simplify inverse Numeral1, a/Numeral1*)
-val inverse_1s = [@{thm inverse_numeral_1}];
-val divide_1s = [@{thm divide_numeral_1}];
-
-(*To perform binary arithmetic.  The "left" rewriting handles patterns
-  created by the Int_Numeral_Base_Simprocs, such as 3 * (5 * x). *)
-val simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym,
-                 add_number_of_left, mult_number_of_left] @
-                arith_simps @ rel_simps;
-
-(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
-  during re-arrangement*)
-val non_add_simps =
-  subtract Thm.eq_thm [add_number_of_left, number_of_add RS sym] simps;
-
-(*To evaluate binary negations of coefficients*)
-val minus_simps = [numeral_m1_eq_minus_1 RS sym, number_of_minus RS sym,
-                   minus_1, minus_0, minus_Pls, minus_Min,
-                   pred_1, pred_0, pred_Pls, pred_Min];
-
-(*To let us treat subtraction as addition*)
-val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
-
-(*To let us treat division as multiplication*)
-val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
-
-(*push the unary minus down: - x * y = x * - y *)
-val minus_mult_eq_1_to_2 =
-    [@{thm minus_mult_left} RS sym, @{thm minus_mult_right}] MRS trans |> standard;
-
-(*to extract again any uncancelled minuses*)
-val minus_from_mult_simps =
-    [@{thm minus_minus}, @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym];
-
-(*combine unary minus with numeric literals, however nested within a product*)
-val mult_minus_simps =
-    [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2];
-
-(*Apply the given rewrite (if present) just once*)
-fun trans_tac NONE      = all_tac
-  | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans));
-
-fun simplify_meta_eq rules =
-  let val ss0 = HOL_basic_ss addeqcongs [eq_cong2] addsimps rules
-  in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end
-
-structure CancelNumeralsCommon =
-  struct
-  val mk_sum            = mk_sum
-  val dest_sum          = dest_sum
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val find_first_coeff  = find_first_coeff []
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
-    diff_simps @ minus_simps @ add_ac
-  val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
-  val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
-
-  val numeral_simp_ss = HOL_ss addsimps add_0s @ simps
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-  val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s)
-  end;
-
-
-structure EqCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
-  val bal_add1 = eq_add_iff1 RS trans
-  val bal_add2 = eq_add_iff2 RS trans
-);
-
-structure LessCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
-  val bal_add1 = less_add_iff1 RS trans
-  val bal_add2 = less_add_iff2 RS trans
-);
-
-structure LeCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
-  val bal_add1 = le_add_iff1 RS trans
-  val bal_add2 = le_add_iff2 RS trans
-);
-
-val cancel_numerals =
-  map Int_Numeral_Base_Simprocs.prep_simproc
-   [("inteq_cancel_numerals",
-     ["(l::'a::number_ring) + m = n",
-      "(l::'a::number_ring) = m + n",
-      "(l::'a::number_ring) - m = n",
-      "(l::'a::number_ring) = m - n",
-      "(l::'a::number_ring) * m = n",
-      "(l::'a::number_ring) = m * n"],
-     K EqCancelNumerals.proc),
-    ("intless_cancel_numerals",
-     ["(l::'a::{ordered_idom,number_ring}) + m < n",
-      "(l::'a::{ordered_idom,number_ring}) < m + n",
-      "(l::'a::{ordered_idom,number_ring}) - m < n",
-      "(l::'a::{ordered_idom,number_ring}) < m - n",
-      "(l::'a::{ordered_idom,number_ring}) * m < n",
-      "(l::'a::{ordered_idom,number_ring}) < m * n"],
-     K LessCancelNumerals.proc),
-    ("intle_cancel_numerals",
-     ["(l::'a::{ordered_idom,number_ring}) + m <= n",
-      "(l::'a::{ordered_idom,number_ring}) <= m + n",
-      "(l::'a::{ordered_idom,number_ring}) - m <= n",
-      "(l::'a::{ordered_idom,number_ring}) <= m - n",
-      "(l::'a::{ordered_idom,number_ring}) * m <= n",
-      "(l::'a::{ordered_idom,number_ring}) <= m * n"],
-     K LeCancelNumerals.proc)];
-
-
-structure CombineNumeralsData =
-  struct
-  type coeff            = IntInf.int
-  val iszero            = (fn x : IntInf.int => x = 0)
-  val add               = IntInf.+
-  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
-  val dest_sum          = dest_sum
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val left_distrib      = combine_common_factor RS trans
-  val prove_conv        = Int_Numeral_Base_Simprocs.prove_conv_nohyps
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
-    diff_simps @ minus_simps @ add_ac
-  val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
-  val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
-
-  val numeral_simp_ss = HOL_ss addsimps add_0s @ simps
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-  val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s)
-  end;
-
-structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
-
-(*Version for fields, where coefficients can be fractions*)
-structure FieldCombineNumeralsData =
-  struct
-  type coeff            = IntInf.int * IntInf.int
-  val iszero            = (fn (p : IntInf.int, q) => p = 0)
-  val add               = add_frac
-  val mk_sum            = long_mk_sum
-  val dest_sum          = dest_sum
-  val mk_coeff          = mk_fcoeff
-  val dest_coeff        = dest_fcoeff 1
-  val left_distrib      = combine_common_factor RS trans
-  val prove_conv        = Int_Numeral_Base_Simprocs.prove_conv_nohyps
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
-    inverse_1s @ divide_simps @ diff_simps @ minus_simps @ add_ac
-  val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
-  val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
-
-  val numeral_simp_ss = HOL_ss addsimps add_0s @ simps @ [@{thm add_frac_eq}]
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-  val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s @ divide_1s)
-  end;
-
-structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData);
-
-val combine_numerals =
-  Int_Numeral_Base_Simprocs.prep_simproc
-    ("int_combine_numerals", 
-     ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], 
-     K CombineNumerals.proc);
-
-val field_combine_numerals =
-  Int_Numeral_Base_Simprocs.prep_simproc
-    ("field_combine_numerals", 
-     ["(i::'a::{number_ring,field,division_by_zero}) + j",
-      "(i::'a::{number_ring,field,division_by_zero}) - j"], 
-     K FieldCombineNumerals.proc);
-
-end;
-
-Addsimprocs Int_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [Int_Numeral_Simprocs.combine_numerals];
-Addsimprocs [Int_Numeral_Simprocs.field_combine_numerals];
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s, by (Simp_tac 1));
-
-test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)";
-
-test "2*u = (u::int)";
-test "(i + j + 12 + (k::int)) - 15 = y";
-test "(i + j + 12 + (k::int)) - 5 = y";
-
-test "y - b < (b::int)";
-test "y - (3*b + c) < (b::int) - 2*c";
-
-test "(2*x - (u*v) + y) - v*3*u = (w::int)";
-test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)";
-test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)";
-test "u*v - (x*u*v + (u*v)*4 + y) = (w::int)";
-
-test "(i + j + 12 + (k::int)) = u + 15 + y";
-test "(i + j*2 + 12 + (k::int)) = j + 5 + y";
-
-test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)";
-
-test "a + -(b+c) + b = (d::int)";
-test "a + -(b+c) - b = (d::int)";
-
-(*negative numerals*)
-test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz";
-test "(i + j + -3 + (k::int)) < u + 5 + y";
-test "(i + j + 3 + (k::int)) < u + -6 + y";
-test "(i + j + -12 + (k::int)) - 15 = y";
-test "(i + j + 12 + (k::int)) - -15 = y";
-test "(i + j + -12 + (k::int)) - -15 = y";
-*)
-
-
-(** Constant folding for multiplication in semirings **)
-
-(*We do not need folding for addition: combine_numerals does the same thing*)
-
-structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA =
-struct
-  val assoc_ss = HOL_ss addsimps mult_ac
-  val eq_reflection = eq_reflection
-end;
-
-structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
-
-val assoc_fold_simproc =
-  Int_Numeral_Base_Simprocs.prep_simproc
-   ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
-    K Semiring_Times_Assoc.proc);
-
-Addsimprocs [assoc_fold_simproc];
-
-
-
-
-(*** decision procedure for linear arithmetic ***)
-
-(*---------------------------------------------------------------------------*)
-(* Linear arithmetic                                                         *)
-(*---------------------------------------------------------------------------*)
-
-(*
-Instantiation of the generic linear arithmetic package for int.
-*)
-
-(* Update parameters of arithmetic prover *)
-local
-
-(* reduce contradictory <= to False *)
-val add_rules =
-    simp_thms @ arith_simps @ rel_simps @ arith_special @
-    [@{thm neg_le_iff_le}, @{thm numeral_0_eq_0}, @{thm numeral_1_eq_1},
-     @{thm minus_zero}, @{thm diff_minus}, @{thm left_minus}, @{thm right_minus},
-     @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_num1}, @{thm mult_1_right},
-     @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym,
-     @{thm minus_add_distrib}, @{thm minus_minus}, @{thm mult_assoc},
-     of_nat_0, of_nat_1, of_nat_Suc, of_nat_add, of_nat_mult,
-     of_int_0, of_int_1, of_int_add, of_int_mult, int_eq_of_nat]
-
-val nat_inj_thms = [zle_int RS iffD2, int_int_eq RS iffD2]
-
-val Int_Numeral_Base_Simprocs = assoc_fold_simproc
-  :: Int_Numeral_Simprocs.combine_numerals
-  :: Int_Numeral_Simprocs.cancel_numerals;
-
-in
-
-val int_arith_setup =
-  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
-   {add_mono_thms = add_mono_thms,
-    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms,
-    inj_thms = nat_inj_thms @ inj_thms,
-    lessD = lessD @ [zless_imp_add1_zle],
-    neqE = neqE,
-    simpset = simpset addsimps add_rules
-                      addsimprocs Int_Numeral_Base_Simprocs
-                      addcongs [if_weak_cong]}) #>
-  arith_inj_const ("Nat.of_nat", HOLogic.natT --> HOLogic.intT) #>
-  arith_inj_const ("IntDef.int", HOLogic.natT --> HOLogic.intT) #>
-  arith_discrete "IntDef.int"
-
-end;
-
-val fast_int_arith_simproc =
-  Simplifier.simproc @{theory}
-  "fast_int_arith" 
-     ["(m::'a::{ordered_idom,number_ring}) < n",
-      "(m::'a::{ordered_idom,number_ring}) <= n",
-      "(m::'a::{ordered_idom,number_ring}) = n"] Fast_Arith.lin_arith_prover;
-
-Addsimprocs [fast_int_arith_simproc];
--- a/src/HOL/Integ/int_factor_simprocs.ML	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,335 +0,0 @@
-(*  Title:      HOL/int_factor_simprocs.ML
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2000  University of Cambridge
-
-Factor cancellation simprocs for the integers (and for fields).
-
-This file can't be combined with int_arith1 because it requires IntDiv.thy.
-*)
-
-
-(*To quote from Provers/Arith/cancel_numeral_factor.ML:
-
-Cancels common coefficients in balanced expressions:
-
-     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
-
-where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
-and d = gcd(m,m') and n=m/d and n'=m'/d.
-*)
-
-val rel_number_of = [eq_number_of_eq, less_number_of_eq_neg, le_number_of_eq];
-
-(** Factor cancellation theorems for integer division (div, not /) **)
-
-Goal "!!k::int. k~=0 ==> (k*m) div (k*n) = (m div n)";
-by (stac @{thm zdiv_zmult_zmult1} 1);
-by Auto_tac;
-qed "int_mult_div_cancel1";
-
-(*For ExtractCommonTermFun, cancelling common factors*)
-Goal "(k*m) div (k*n) = (if k = (0::int) then 0 else m div n)";
-by (simp_tac (simpset() addsimps [int_mult_div_cancel1]) 1);
-qed "int_mult_div_cancel_disj";
-
-
-local
-  open Int_Numeral_Simprocs
-in
-
-structure CancelNumeralFactorCommon =
-  struct
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = HOL_ss addsimps minus_from_mult_simps @ mult_1s
-  val norm_ss2 = HOL_ss addsimps simps @ mult_minus_simps
-  val norm_ss3 = HOL_ss addsimps mult_ac
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
-
-  val numeral_simp_ss = HOL_ss addsimps rel_number_of @ simps
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-  val simplify_meta_eq = Int_Numeral_Simprocs.simplify_meta_eq
-    [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left},
-      @{thm mult_zero_right}, @{thm mult_num1}, @{thm mult_1_right}];
-  end
-
-(*Version for integer division*)
-structure IntDivCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT
-  val cancel = int_mult_div_cancel1 RS trans
-  val neg_exchanges = false
-)
-
-(*Version for fields*)
-structure DivideCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name HOL.divide}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT
-  val cancel = @{thm mult_divide_cancel_left} RS trans
-  val neg_exchanges = false
-)
-
-structure EqCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
-  val cancel = @{thm mult_cancel_left} RS trans
-  val neg_exchanges = false
-)
-
-structure LessCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
-  val cancel = @{thm mult_less_cancel_left} RS trans
-  val neg_exchanges = true
-)
-
-structure LeCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
-  val cancel = @{thm mult_le_cancel_left} RS trans
-  val neg_exchanges = true
-)
-
-val cancel_numeral_factors =
-  map Int_Numeral_Base_Simprocs.prep_simproc
-   [("ring_eq_cancel_numeral_factor",
-     ["(l::'a::{idom,number_ring}) * m = n",
-      "(l::'a::{idom,number_ring}) = m * n"],
-     K EqCancelNumeralFactor.proc),
-    ("ring_less_cancel_numeral_factor",
-     ["(l::'a::{ordered_idom,number_ring}) * m < n",
-      "(l::'a::{ordered_idom,number_ring}) < m * n"],
-     K LessCancelNumeralFactor.proc),
-    ("ring_le_cancel_numeral_factor",
-     ["(l::'a::{ordered_idom,number_ring}) * m <= n",
-      "(l::'a::{ordered_idom,number_ring}) <= m * n"],
-     K LeCancelNumeralFactor.proc),
-    ("int_div_cancel_numeral_factors",
-     ["((l::int) * m) div n", "(l::int) div (m * n)"],
-     K IntDivCancelNumeralFactor.proc),
-    ("divide_cancel_numeral_factor",
-     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
-      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
-      "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
-     K DivideCancelNumeralFactor.proc)];
-
-(* referenced by rat_arith.ML *)
-val field_cancel_numeral_factors =
-  map Int_Numeral_Base_Simprocs.prep_simproc
-   [("field_eq_cancel_numeral_factor",
-     ["(l::'a::{field,number_ring}) * m = n",
-      "(l::'a::{field,number_ring}) = m * n"],
-     K EqCancelNumeralFactor.proc),
-    ("field_cancel_numeral_factor",
-     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
-      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
-      "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
-     K DivideCancelNumeralFactor.proc)]
-
-end;
-
-Addsimprocs cancel_numeral_factors;
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Simp_tac 1));
-
-test "9*x = 12 * (y::int)";
-test "(9*x) div (12 * (y::int)) = z";
-test "9*x < 12 * (y::int)";
-test "9*x <= 12 * (y::int)";
-
-test "-99*x = 132 * (y::int)";
-test "(-99*x) div (132 * (y::int)) = z";
-test "-99*x < 132 * (y::int)";
-test "-99*x <= 132 * (y::int)";
-
-test "999*x = -396 * (y::int)";
-test "(999*x) div (-396 * (y::int)) = z";
-test "999*x < -396 * (y::int)";
-test "999*x <= -396 * (y::int)";
-
-test "-99*x = -81 * (y::int)";
-test "(-99*x) div (-81 * (y::int)) = z";
-test "-99*x <= -81 * (y::int)";
-test "-99*x < -81 * (y::int)";
-
-test "-2 * x = -1 * (y::int)";
-test "-2 * x = -(y::int)";
-test "(-2 * x) div (-1 * (y::int)) = z";
-test "-2 * x < -(y::int)";
-test "-2 * x <= -1 * (y::int)";
-test "-x < -23 * (y::int)";
-test "-x <= -23 * (y::int)";
-*)
-
-(*And the same examples for fields such as rat or real:
-test "0 <= (y::rat) * -2";
-test "9*x = 12 * (y::rat)";
-test "(9*x) / (12 * (y::rat)) = z";
-test "9*x < 12 * (y::rat)";
-test "9*x <= 12 * (y::rat)";
-
-test "-99*x = 132 * (y::rat)";
-test "(-99*x) / (132 * (y::rat)) = z";
-test "-99*x < 132 * (y::rat)";
-test "-99*x <= 132 * (y::rat)";
-
-test "999*x = -396 * (y::rat)";
-test "(999*x) / (-396 * (y::rat)) = z";
-test "999*x < -396 * (y::rat)";
-test "999*x <= -396 * (y::rat)";
-
-test  "(- ((2::rat) * x) <= 2 * y)";
-test "-99*x = -81 * (y::rat)";
-test "(-99*x) / (-81 * (y::rat)) = z";
-test "-99*x <= -81 * (y::rat)";
-test "-99*x < -81 * (y::rat)";
-
-test "-2 * x = -1 * (y::rat)";
-test "-2 * x = -(y::rat)";
-test "(-2 * x) / (-1 * (y::rat)) = z";
-test "-2 * x < -(y::rat)";
-test "-2 * x <= -1 * (y::rat)";
-test "-x < -23 * (y::rat)";
-test "-x <= -23 * (y::rat)";
-*)
-
-
-(** Declarations for ExtractCommonTerm **)
-
-local
-  open Int_Numeral_Simprocs
-in
-
-(*Find first term that matches u*)
-fun find_first_t past u []         = raise TERM ("find_first_t", [])
-  | find_first_t past u (t::terms) =
-        if u aconv t then (rev past @ terms)
-        else find_first_t (t::past) u terms
-        handle TERM _ => find_first_t (t::past) u terms;
-
-(** Final simplification for the CancelFactor simprocs **)
-val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq  
-  [@{thm mult_1_left}, mult_1_right, @{thm zdiv_1}, numeral_1_eq_1];
-
-fun cancel_simplify_meta_eq cancel_th ss th =
-    simplify_one ss (([th, cancel_th]) MRS trans);
-
-(*At present, long_mk_prod creates Numeral1, so this requires the axclass
-  number_ring*)
-structure CancelFactorCommon =
-  struct
-  val mk_sum            = long_mk_prod
-  val dest_sum          = dest_prod
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff
-  val find_first        = find_first_t []
-  val trans_tac         = fn _ => trans_tac
-  val norm_ss = HOL_ss addsimps mult_1s @ mult_ac
-  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
-  end;
-
-(*mult_cancel_left requires a ring with no zero divisors.*)
-structure EqCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_cancel_left}
-);
-
-(*int_mult_div_cancel_disj is for integer division (div).*)
-structure IntDivCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT
-  val simplify_meta_eq  = cancel_simplify_meta_eq int_mult_div_cancel_disj
-);
-
-(*Version for all fields, including unordered ones (type complex).*)
-structure DivideCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name HOL.divide}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT
-  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_divide_cancel_eq_if}
-);
-
-(*The number_ring class is necessary because the simprocs refer to the
-  binary number 1.  FIXME: probably they could use 1 instead.*)
-val cancel_factors =
-  map Int_Numeral_Base_Simprocs.prep_simproc
-   [("ring_eq_cancel_factor",
-     ["(l::'a::{idom,number_ring}) * m = n",
-      "(l::'a::{idom,number_ring}) = m * n"],
-    K EqCancelFactor.proc),
-    ("int_div_cancel_factor",
-     ["((l::int) * m) div n", "(l::int) div (m * n)"],
-     K IntDivCancelFactor.proc),
-    ("divide_cancel_factor",
-     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
-      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)"],
-     K DivideCancelFactor.proc)];
-
-end;
-
-Addsimprocs cancel_factors;
-
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Asm_simp_tac 1));
-
-test "x*k = k*(y::int)";
-test "k = k*(y::int)";
-test "a*(b*c) = (b::int)";
-test "a*(b*c) = d*(b::int)*(x*a)";
-
-test "(x*k) div (k*(y::int)) = (uu::int)";
-test "(k) div (k*(y::int)) = (uu::int)";
-test "(a*(b*c)) div ((b::int)) = (uu::int)";
-test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)";
-*)
-
-(*And the same examples for fields such as rat or real:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Asm_simp_tac 1));
-
-test "x*k = k*(y::rat)";
-test "k = k*(y::rat)";
-test "a*(b*c) = (b::rat)";
-test "a*(b*c) = d*(b::rat)*(x*a)";
-
-
-test "(x*k) / (k*(y::rat)) = (uu::rat)";
-test "(k) / (k*(y::rat)) = (uu::rat)";
-test "(a*(b*c)) / ((b::rat)) = (uu::rat)";
-test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)";
-
-(*FIXME: what do we do about this?*)
-test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z";
-*)
--- a/src/HOL/Integ/nat_simprocs.ML	Thu May 31 18:16:51 2007 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,550 +0,0 @@
-(*  Title:      HOL/nat_simprocs.ML
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   2000  University of Cambridge
-
-Simprocs for nat numerals.
-*)
-
-structure Nat_Numeral_Simprocs =
-struct
-
-(*Maps n to #n for n = 0, 1, 2*)
-val numeral_syms =
-       [nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym, numeral_2_eq_2 RS sym];
-val numeral_sym_ss = HOL_ss addsimps numeral_syms;
-
-fun rename_numerals th =
-    simplify numeral_sym_ss (Thm.transfer (the_context ()) th);
-
-(*Utilities*)
-
-fun mk_number n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_numeral n;
-fun dest_number t = IntInf.max (0, snd (HOLogic.dest_number t));
-
-fun find_first_numeral past (t::terms) =
-        ((dest_number t, t, rev past @ terms)
-         handle TERM _ => find_first_numeral (t::past) terms)
-  | find_first_numeral past [] = raise TERM("find_first_numeral", []);
-
-val zero = mk_number 0;
-val mk_plus = HOLogic.mk_binop @{const_name HOL.plus};
-
-(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
-fun mk_sum []        = zero
-  | mk_sum [t,u]     = mk_plus (t, u)
-  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
-
-(*this version ALWAYS includes a trailing zero*)
-fun long_mk_sum []        = HOLogic.zero
-  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
-
-val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT;
-
-
-(** Other simproc items **)
-
-val trans_tac = Int_Numeral_Simprocs.trans_tac;
-
-val bin_simps =
-     [nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym,
-      add_nat_number_of, nat_number_of_add_left, 
-      diff_nat_number_of, le_number_of_eq_not_less,
-      mult_nat_number_of, nat_number_of_mult_left, 
-      less_nat_number_of, 
-      @{thm Let_number_of}, nat_number_of] @
-     arith_simps @ rel_simps;
-
-fun prep_simproc (name, pats, proc) =
-  Simplifier.simproc (the_context ()) name pats proc;
-
-
-(*** CancelNumerals simprocs ***)
-
-val one = mk_number 1;
-val mk_times = HOLogic.mk_binop @{const_name HOL.times};
-
-fun mk_prod [] = one
-  | mk_prod [t] = t
-  | mk_prod (t :: ts) = if t = one then mk_prod ts
-                        else mk_times (t, mk_prod ts);
-
-val dest_times = HOLogic.dest_bin @{const_name HOL.times} HOLogic.natT;
-
-fun dest_prod t =
-      let val (t,u) = dest_times t
-      in  dest_prod t @ dest_prod u  end
-      handle TERM _ => [t];
-
-(*DON'T do the obvious simplifications; that would create special cases*)
-fun mk_coeff (k,t) = mk_times (mk_number k, t);
-
-(*Express t as a product of (possibly) a numeral with other factors, sorted*)
-fun dest_coeff t =
-    let val ts = sort Term.term_ord (dest_prod t)
-        val (n, _, ts') = find_first_numeral [] ts
-                          handle TERM _ => (1, one, ts)
-    in (n, mk_prod ts') end;
-
-(*Find first coefficient-term THAT MATCHES u*)
-fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
-  | find_first_coeff past u (t::terms) =
-        let val (n,u') = dest_coeff t
-        in  if u aconv u' then (n, rev past @ terms)
-                          else find_first_coeff (t::past) u terms
-        end
-        handle TERM _ => find_first_coeff (t::past) u terms;
-
-
-(*Split up a sum into the list of its constituent terms, on the way removing any
-  Sucs and counting them.*)
-fun dest_Suc_sum (Const ("Suc", _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts))
-  | dest_Suc_sum (t, (k,ts)) = 
-      let val (t1,t2) = dest_plus t
-      in  dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts)))  end
-      handle TERM _ => (k, t::ts);
-
-(*Code for testing whether numerals are already used in the goal*)
-fun is_numeral (Const(@{const_name Numeral.number_of}, _) $ w) = true
-  | is_numeral _ = false;
-
-fun prod_has_numeral t = exists is_numeral (dest_prod t);
-
-(*The Sucs found in the term are converted to a binary numeral. If relaxed is false,
-  an exception is raised unless the original expression contains at least one
-  numeral in a coefficient position.  This prevents nat_combine_numerals from 
-  introducing numerals to goals.*)
-fun dest_Sucs_sum relaxed t = 
-  let val (k,ts) = dest_Suc_sum (t,(0,[]))
-  in
-     if relaxed orelse exists prod_has_numeral ts then 
-       if k=0 then ts
-       else mk_number (IntInf.fromInt k) :: ts
-     else raise TERM("Nat_Numeral_Simprocs.dest_Sucs_sum", [t])
-  end;
-
-
-(*Simplify 1*n and n*1 to n*)
-val add_0s  = map rename_numerals [add_0, add_0_right];
-val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}];
-
-(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*)
-
-(*And these help the simproc return False when appropriate, which helps
-  the arith prover.*)
-val contra_rules = [add_Suc, add_Suc_right, Zero_not_Suc, Suc_not_Zero,
-                    le_0_eq];
-
-val simplify_meta_eq =
-    Int_Numeral_Simprocs.simplify_meta_eq
-        ([nat_numeral_0_eq_0, numeral_1_eq_Suc_0, add_0, add_0_right,
-          mult_0, mult_0_right, mult_1, mult_1_right] @ contra_rules);
-
-
-(*Like HOL_ss but with an ordering that brings numerals to the front
-  under AC-rewriting.*)
-val num_ss = Int_Numeral_Simprocs.num_ss;
-
-(*** Applying CancelNumeralsFun ***)
-
-structure CancelNumeralsCommon =
-  struct
-  val mk_sum            = (fn T:typ => mk_sum)
-  val dest_sum          = dest_Sucs_sum true
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff
-  val find_first_coeff  = find_first_coeff []
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
-    [Suc_eq_add_numeral_1_left] @ add_ac
-  val norm_ss2 = num_ss addsimps bin_simps @ add_ac @ mult_ac
-  fun norm_tac ss = 
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-
-  val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps;
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss));
-  val simplify_meta_eq  = simplify_meta_eq
-  end;
-
-
-structure EqCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val bal_add1 = nat_eq_add_iff1 RS trans
-  val bal_add2 = nat_eq_add_iff2 RS trans
-);
-
-structure LessCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
-  val bal_add1 = nat_less_add_iff1 RS trans
-  val bal_add2 = nat_less_add_iff2 RS trans
-);
-
-structure LeCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
-  val bal_add1 = nat_le_add_iff1 RS trans
-  val bal_add2 = nat_le_add_iff2 RS trans
-);
-
-structure DiffCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name HOL.minus}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT
-  val bal_add1 = nat_diff_add_eq1 RS trans
-  val bal_add2 = nat_diff_add_eq2 RS trans
-);
-
-
-val cancel_numerals =
-  map prep_simproc
-   [("nateq_cancel_numerals",
-     ["(l::nat) + m = n", "(l::nat) = m + n",
-      "(l::nat) * m = n", "(l::nat) = m * n",
-      "Suc m = n", "m = Suc n"],
-     K EqCancelNumerals.proc),
-    ("natless_cancel_numerals",
-     ["(l::nat) + m < n", "(l::nat) < m + n",
-      "(l::nat) * m < n", "(l::nat) < m * n",
-      "Suc m < n", "m < Suc n"],
-     K LessCancelNumerals.proc),
-    ("natle_cancel_numerals",
-     ["(l::nat) + m <= n", "(l::nat) <= m + n",
-      "(l::nat) * m <= n", "(l::nat) <= m * n",
-      "Suc m <= n", "m <= Suc n"],
-     K LeCancelNumerals.proc),
-    ("natdiff_cancel_numerals",
-     ["((l::nat) + m) - n", "(l::nat) - (m + n)",
-      "(l::nat) * m - n", "(l::nat) - m * n",
-      "Suc m - n", "m - Suc n"],
-     K DiffCancelNumerals.proc)];
-
-
-(*** Applying CombineNumeralsFun ***)
-
-structure CombineNumeralsData =
-  struct
-  type coeff            = IntInf.int
-  val iszero            = (fn x : IntInf.int => x = 0)
-  val add               = IntInf.+
-  val mk_sum            = (fn T:typ => long_mk_sum)  (*to work for 2*x + 3*x *)
-  val dest_sum          = dest_Sucs_sum false
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff
-  val left_distrib      = left_add_mult_distrib RS trans
-  val prove_conv        = Int_Numeral_Base_Simprocs.prove_conv_nohyps
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [Suc_eq_add_numeral_1] @ add_ac
-  val norm_ss2 = num_ss addsimps bin_simps @ add_ac @ mult_ac
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-
-  val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps;
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-  val simplify_meta_eq  = simplify_meta_eq
-  end;
-
-structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
-
-val combine_numerals =
-  prep_simproc ("nat_combine_numerals", ["(i::nat) + j", "Suc (i + j)"], K CombineNumerals.proc);
-
-
-(*** Applying CancelNumeralFactorFun ***)
-
-structure CancelNumeralFactorCommon =
-  struct
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff
-  val trans_tac         = fn _ => trans_tac
-
-  val norm_ss1 = num_ss addsimps
-    numeral_syms @ add_0s @ mult_1s @ [Suc_eq_add_numeral_1_left] @ add_ac
-  val norm_ss2 = num_ss addsimps bin_simps @ add_ac @ mult_ac
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
-    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
-
-  val numeral_simp_ss = HOL_ss addsimps bin_simps
-  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
-  val simplify_meta_eq  = simplify_meta_eq
-  end
-
-structure DivCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
-  val cancel = nat_mult_div_cancel1 RS trans
-  val neg_exchanges = false
-)
-
-structure EqCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val cancel = nat_mult_eq_cancel1 RS trans
-  val neg_exchanges = false
-)
-
-structure LessCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
-  val cancel = nat_mult_less_cancel1 RS trans
-  val neg_exchanges = true
-)
-
-structure LeCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
-  val cancel = nat_mult_le_cancel1 RS trans
-  val neg_exchanges = true
-)
-
-val cancel_numeral_factors =
-  map prep_simproc
-   [("nateq_cancel_numeral_factors",
-     ["(l::nat) * m = n", "(l::nat) = m * n"],
-     K EqCancelNumeralFactor.proc),
-    ("natless_cancel_numeral_factors",
-     ["(l::nat) * m < n", "(l::nat) < m * n"],
-     K LessCancelNumeralFactor.proc),
-    ("natle_cancel_numeral_factors",
-     ["(l::nat) * m <= n", "(l::nat) <= m * n"],
-     K LeCancelNumeralFactor.proc),
-    ("natdiv_cancel_numeral_factors",
-     ["((l::nat) * m) div n", "(l::nat) div (m * n)"],
-     K DivCancelNumeralFactor.proc)];
-
-
-
-(*** Applying ExtractCommonTermFun ***)
-
-(*this version ALWAYS includes a trailing one*)
-fun long_mk_prod []        = one
-  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
-
-(*Find first term that matches u*)
-fun find_first_t past u []         = raise TERM("find_first_t", [])
-  | find_first_t past u (t::terms) =
-        if u aconv t then (rev past @ terms)
-        else find_first_t (t::past) u terms
-        handle TERM _ => find_first_t (t::past) u terms;
-
-(** Final simplification for the CancelFactor simprocs **)
-val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq  
-  [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}];
-
-fun cancel_simplify_meta_eq cancel_th ss th =
-    simplify_one ss (([th, cancel_th]) MRS trans);
-
-structure CancelFactorCommon =
-  struct
-  val mk_sum            = (fn T:typ => long_mk_prod)
-  val dest_sum          = dest_prod
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff
-  val find_first        = find_first_t []
-  val trans_tac         = fn _ => trans_tac
-  val norm_ss = HOL_ss addsimps mult_1s @ mult_ac
-  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
-  end;
-
-structure EqCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_eq_cancel_disj
-);
-
-structure LessCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_less_cancel_disj
-);
-
-structure LeCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_le_cancel_disj
-);
-
-structure DivideCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
-  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_div_cancel_disj
-);
-
-val cancel_factor =
-  map prep_simproc
-   [("nat_eq_cancel_factor",
-     ["(l::nat) * m = n", "(l::nat) = m * n"],
-     K EqCancelFactor.proc),
-    ("nat_less_cancel_factor",
-     ["(l::nat) * m < n", "(l::nat) < m * n"],
-     K LessCancelFactor.proc),
-    ("nat_le_cancel_factor",
-     ["(l::nat) * m <= n", "(l::nat) <= m * n"],
-     K LeCancelFactor.proc),
-    ("nat_divide_cancel_factor",
-     ["((l::nat) * m) div n", "(l::nat) div (m * n)"],
-     K DivideCancelFactor.proc)];
-
-end;
-
-
-Addsimprocs Nat_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [Nat_Numeral_Simprocs.combine_numerals];
-Addsimprocs Nat_Numeral_Simprocs.cancel_numeral_factors;
-Addsimprocs Nat_Numeral_Simprocs.cancel_factor;
-
-
-(*examples:
-print_depth 22;
-set timing;
-set trace_simp;
-fun test s = (Goal s; by (Simp_tac 1));
-
-(*cancel_numerals*)
-test "l +( 2) + (2) + 2 + (l + 2) + (oo  + 2) = (uu::nat)";
-test "(2*length xs < 2*length xs + j)";
-test "(2*length xs < length xs * 2 + j)";
-test "2*u = (u::nat)";
-test "2*u = Suc (u)";
-test "(i + j + 12 + (k::nat)) - 15 = y";
-test "(i + j + 12 + (k::nat)) - 5 = y";
-test "Suc u - 2 = y";
-test "Suc (Suc (Suc u)) - 2 = y";
-test "(i + j + 2 + (k::nat)) - 1 = y";
-test "(i + j + 1 + (k::nat)) - 2 = y";
-
-test "(2*x + (u*v) + y) - v*3*u = (w::nat)";
-test "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = (w::nat)";
-test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::nat)";
-test "Suc (Suc (2*x*u*v + u*4 + y)) - u = w";
-test "Suc ((u*v)*4) - v*3*u = w";
-test "Suc (Suc ((u*v)*3)) - v*3*u = w";
-
-test "(i + j + 12 + (k::nat)) = u + 15 + y";
-test "(i + j + 32 + (k::nat)) - (u + 15 + y) = zz";
-test "(i + j + 12 + (k::nat)) = u + 5 + y";
-(*Suc*)
-test "(i + j + 12 + k) = Suc (u + y)";
-test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + 41 + k)";
-test "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))";
-test "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v";
-test "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))";
-test "2*y + 3*z + 2*u = Suc (u)";
-test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)";
-test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::nat)";
-test "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)";
-test "(2*n*m) < (3*(m*n)) + (u::nat)";
-
-test "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) <= Suc 0)";
- 
-test "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) <= length l1";
-
-test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) <= length (compT P E A ST mxr e))";
-
-test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un \<A> e) ST mxr c))))))) <= length (compT P E A ST mxr e))";
-
-
-(*negative numerals: FAIL*)
-test "(i + j + -23 + (k::nat)) < u + 15 + y";
-test "(i + j + 3 + (k::nat)) < u + -15 + y";
-test "(i + j + -12 + (k::nat)) - 15 = y";
-test "(i + j + 12 + (k::nat)) - -15 = y";
-test "(i + j + -12 + (k::nat)) - -15 = y";
-
-(*combine_numerals*)
-test "k + 3*k = (u::nat)";
-test "Suc (i + 3) = u";
-test "Suc (i + j + 3 + k) = u";
-test "k + j + 3*k + j = (u::nat)";
-test "Suc (j*i + i + k + 5 + 3*k + i*j*4) = (u::nat)";
-test "(2*n*m) + (3*(m*n)) = (u::nat)";
-(*negative numerals: FAIL*)
-test "Suc (i + j + -3 + k) = u";
-
-(*cancel_numeral_factors*)
-test "9*x = 12 * (y::nat)";
-test "(9*x) div (12 * (y::nat)) = z";
-test "9*x < 12 * (y::nat)";
-test "9*x <= 12 * (y::nat)";
-
-(*cancel_factor*)
-test "x*k = k*(y::nat)";
-test "k = k*(y::nat)";
-test "a*(b*c) = (b::nat)";
-test "a*(b*c) = d*(b::nat)*(x*a)";
-
-test "x*k < k*(y::nat)";
-test "k < k*(y::nat)";
-test "a*(b*c) < (b::nat)";
-test "a*(b*c) < d*(b::nat)*(x*a)";
-
-test "x*k <= k*(y::nat)";
-test "k <= k*(y::nat)";
-test "a*(b*c) <= (b::nat)";
-test "a*(b*c) <= d*(b::nat)*(x*a)";
-
-test "(x*k) div (k*(y::nat)) = (uu::nat)";
-test "(k) div (k*(y::nat)) = (uu::nat)";
-test "(a*(b*c)) div ((b::nat)) = (uu::nat)";
-test "(a*(b*c)) div (d*(b::nat)*(x*a)) = (uu::nat)";
-*)
-
-
-(*** Prepare linear arithmetic for nat numerals ***)
-
-local
-
-(* reduce contradictory <= to False *)
-val add_rules =
-  [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, nat_0, nat_1,
-   add_nat_number_of, diff_nat_number_of, mult_nat_number_of,
-   eq_nat_number_of, less_nat_number_of, le_number_of_eq_not_less,
-   le_Suc_number_of,le_number_of_Suc,
-   less_Suc_number_of,less_number_of_Suc,
-   Suc_eq_number_of,eq_number_of_Suc,
-   mult_Suc, mult_Suc_right,
-   eq_number_of_0, eq_0_number_of, less_0_number_of,
-   of_int_number_of_eq, of_nat_number_of_eq, nat_number_of, if_True, if_False];
-
-val simprocs = Nat_Numeral_Simprocs.combine_numerals
-  :: Nat_Numeral_Simprocs.cancel_numerals;
-
-in
-
-val nat_simprocs_setup =
-  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
-   {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
-    inj_thms = inj_thms, lessD = lessD, neqE = neqE,
-    simpset = simpset addsimps add_rules
-                      addsimprocs simprocs});
-
-end;
--- a/src/HOL/IsaMakefile	Thu May 31 18:16:51 2007 +0200
+++ b/src/HOL/IsaMakefile	Thu May 31 18:16:52 2007 +0200
@@ -82,26 +82,23 @@
   $(SRC)/Provers/quasi.ML $(SRC)/Provers/splitter.ML				\
   $(SRC)/Provers/trancl.ML $(SRC)/Pure/General/int.ML				\
   $(SRC)/Pure/General/rat.ML Tools/TFL/casesplit.ML				\
-  Tools/TFL/dcterm.ML Tools/TFL/post.ML Tools/TFL/rules.ML			\
-  Tools/TFL/tfl.ML Tools/TFL/thms.ML Tools/TFL/thry.ML				\
-  Tools/TFL/usyntax.ML Tools/TFL/utils.ML ATP_Linkup.thy			\
-  Accessible_Part.thy Code_Generator.thy Datatype.thy Divides.thy		\
-  Equiv_Relations.thy Extraction.thy Finite_Set.thy FixedPoint.thy		\
-  Fun.thy FunDef.thy HOL.thy Hilbert_Choice.thy Inductive.thy			\
-  Integ/IntArith.thy Integ/IntDef.thy Integ/IntDiv.thy Integ/NatBin.thy		\
-  Integ/NatSimprocs.thy Integ/Numeral.thy Integ/int_arith1.ML			\
-  Integ/int_factor_simprocs.ML Integ/nat_simprocs.ML Lattices.thy		\
-  List.thy Main.thy Map.thy Nat.ML Nat.thy OrderedGroup.thy			\
-  Orderings.thy Power.thy PreList.thy Predicate.thy Presburger.thy		\
-  Product_Type.thy ROOT.ML Recdef.thy Record.thy Refute.thy Relation.thy	\
-  Relation_Power.thy Ring_and_Field.thy SAT.thy Set.thy SetInterval.thy		\
-  Sum_Type.thy Tools/ATP/reduce_axiomsN.ML Tools/ATP/watcher.ML			\
-  Tools/Presburger/cooper_dec.ML Tools/Presburger/cooper_proof.ML		\
-  Tools/Presburger/presburger.ML Tools/Presburger/qelim.ML			\
-  Tools/Presburger/reflected_cooper.ML						\
-  Tools/Presburger/reflected_presburger.ML Tools/cnf_funcs.ML			\
-  Tools/datatype_abs_proofs.ML Tools/datatype_aux.ML				\
-  Tools/datatype_case.ML Tools/datatype_codegen.ML				\
+  ATP_Linkup.thy Accessible_Part.thy Code_Generator.thy Datatype.thy		\
+  Divides.thy Equiv_Relations.thy Extraction.thy Finite_Set.thy			\
+  FixedPoint.thy Fun.thy FunDef.thy HOL.thy Hilbert_Choice.thy			\
+  Inductive.thy IntArith.thy IntDef.thy IntDiv.thy Lattices.thy List.thy	\
+  Main.thy Map.thy Nat.ML Nat.thy NatBin.thy NatSimprocs.thy Numeral.thy	\
+  OrderedGroup.thy Orderings.thy Power.thy PreList.thy Predicate.thy		\
+  Presburger.thy Product_Type.thy ROOT.ML Recdef.thy Record.thy			\
+  Refute.thy Relation.thy Relation_Power.thy Ring_and_Field.thy SAT.thy		\
+  Set.thy SetInterval.thy Sum_Type.thy Tools/ATP/reduce_axiomsN.ML		\
+  Tools/ATP/watcher.ML Tools/Presburger/cooper_dec.ML				\
+  Tools/Presburger/cooper_proof.ML Tools/Presburger/presburger.ML		\
+  Tools/Presburger/qelim.ML Tools/Presburger/reflected_cooper.ML		\
+  Tools/Presburger/reflected_presburger.ML Tools/TFL/dcterm.ML			\
+  Tools/TFL/post.ML Tools/TFL/rules.ML Tools/TFL/tfl.ML				\
+  Tools/TFL/thms.ML Tools/TFL/thry.ML Tools/TFL/usyntax.ML			\
+  Tools/TFL/utils.ML Tools/cnf_funcs.ML Tools/datatype_abs_proofs.ML		\
+  Tools/datatype_aux.ML Tools/datatype_case.ML Tools/datatype_codegen.ML	\
   Tools/datatype_hooks.ML Tools/datatype_package.ML				\
   Tools/datatype_prop.ML Tools/datatype_realizer.ML				\
   Tools/datatype_rep_proofs.ML Tools/function_package/auto_term.ML		\
@@ -130,7 +127,7 @@
   Tools/typedef_codegen.ML Tools/typedef_package.ML				\
   Transitive_Closure.thy Typedef.thy Wellfounded_Recursion.thy			\
   Wellfounded_Relations.thy arith_data.ML document/root.tex hologic.ML		\
-  simpdata.ML
+  int_arith1.ML int_factor_simprocs.ML nat_simprocs.ML simpdata.ML
 	@$(ISATOOL) usedir $(HOL_USEDIR_OPTIONS) -b -g true $(OUT)/Pure HOL
 
 
@@ -422,7 +419,7 @@
   Auth/Guard/Proto.thy Auth/Guard/Guard_Yahalom.thy\
   Auth/Smartcard/EventSC.thy Auth/Smartcard/ShoupRubinBella.thy\
   Auth/Smartcard/ShoupRubin.thy Auth/Smartcard/Smartcard.thy\
-  Auth/document/root.tex 
+  Auth/document/root.tex
 	@$(ISATOOL) usedir -g true $(OUT)/HOL Auth
 
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NatBin.thy	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,903 @@
+(*  Title:      HOL/NatBin.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+*)
+
+header {* Binary arithmetic for the natural numbers *}
+
+theory NatBin
+imports IntDiv
+begin
+
+text {*
+  Arithmetic for naturals is reduced to that for the non-negative integers.
+*}
+
+instance nat :: number
+  nat_number_of_def [code inline]: "number_of v == nat (number_of (v\<Colon>int))" ..
+
+abbreviation (xsymbols)
+  square :: "'a::power => 'a"  ("(_\<twosuperior>)" [1000] 999) where
+  "x\<twosuperior> == x^2"
+
+notation (latex output)
+  square  ("(_\<twosuperior>)" [1000] 999)
+
+notation (HTML output)
+  square  ("(_\<twosuperior>)" [1000] 999)
+
+
+subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*}
+
+declare nat_0 [simp] nat_1 [simp]
+
+lemma nat_number_of [simp]: "nat (number_of w) = number_of w"
+by (simp add: nat_number_of_def)
+
+lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)"
+by (simp add: nat_number_of_def)
+
+lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
+by (simp add: nat_1 nat_number_of_def)
+
+lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
+by (simp add: nat_numeral_1_eq_1)
+
+lemma numeral_2_eq_2: "2 = Suc (Suc 0)"
+apply (unfold nat_number_of_def)
+apply (rule nat_2)
+done
+
+
+text{*Distributive laws for type @{text nat}.  The others are in theory
+   @{text IntArith}, but these require div and mod to be defined for type
+   "int".  They also need some of the lemmas proved above.*}
+
+lemma nat_div_distrib: "(0::int) <= z ==> nat (z div z') = nat z div nat z'"
+apply (case_tac "0 <= z'")
+apply (auto simp add: div_nonneg_neg_le0 DIVISION_BY_ZERO_DIV)
+apply (case_tac "z' = 0", simp add: DIVISION_BY_ZERO)
+apply (auto elim!: nonneg_eq_int)
+apply (rename_tac m m')
+apply (subgoal_tac "0 <= int m div int m'")
+ prefer 2 apply (simp add: nat_numeral_0_eq_0 pos_imp_zdiv_nonneg_iff) 
+apply (rule inj_int [THEN injD], simp)
+apply (rule_tac r = "int (m mod m') " in quorem_div)
+ prefer 2 apply force
+apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int 
+                 zmult_int)
+done
+
+(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*)
+lemma nat_mod_distrib:
+     "[| (0::int) <= z;  0 <= z' |] ==> nat (z mod z') = nat z mod nat z'"
+apply (case_tac "z' = 0", simp add: DIVISION_BY_ZERO)
+apply (auto elim!: nonneg_eq_int)
+apply (rename_tac m m')
+apply (subgoal_tac "0 <= int m mod int m'")
+ prefer 2 apply (simp add: nat_less_iff nat_numeral_0_eq_0 pos_mod_sign) 
+apply (rule inj_int [THEN injD], simp)
+apply (rule_tac q = "int (m div m') " in quorem_mod)
+ prefer 2 apply force
+apply (simp add: nat_less_iff [symmetric] quorem_def nat_numeral_0_eq_0 zadd_int zmult_int)
+done
+
+text{*Suggested by Matthias Daum*}
+lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)"
+apply (subgoal_tac "nat x div nat k < nat x")
+ apply (simp (asm_lr) add: nat_div_distrib [symmetric])
+apply (rule Divides.div_less_dividend, simp_all) 
+done
+
+subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*}
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+lemma int_nat_number_of [simp]:
+     "int (number_of v :: nat) =  
+         (if neg (number_of v :: int) then 0  
+          else (number_of v :: int))"
+by (simp del: nat_number_of
+	 add: neg_nat nat_number_of_def not_neg_nat add_assoc)
+
+
+subsubsection{*Successor *}
+
+lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)"
+apply (rule sym)
+apply (simp add: nat_eq_iff int_Suc)
+done
+
+lemma Suc_nat_number_of_add:
+     "Suc (number_of v + n) =  
+        (if neg (number_of v :: int) then 1+n else number_of (Numeral.succ v) + n)" 
+by (simp del: nat_number_of 
+         add: nat_number_of_def neg_nat
+              Suc_nat_eq_nat_zadd1 number_of_succ) 
+
+lemma Suc_nat_number_of [simp]:
+     "Suc (number_of v) =  
+        (if neg (number_of v :: int) then 1 else number_of (Numeral.succ v))"
+apply (cut_tac n = 0 in Suc_nat_number_of_add)
+apply (simp cong del: if_weak_cong)
+done
+
+
+subsubsection{*Addition *}
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+lemma add_nat_number_of [simp]:
+     "(number_of v :: nat) + number_of v' =  
+         (if neg (number_of v :: int) then number_of v'  
+          else if neg (number_of v' :: int) then number_of v  
+          else number_of (v + v'))"
+by (force dest!: neg_nat
+          simp del: nat_number_of
+          simp add: nat_number_of_def nat_add_distrib [symmetric]) 
+
+
+subsubsection{*Subtraction *}
+
+lemma diff_nat_eq_if:
+     "nat z - nat z' =  
+        (if neg z' then nat z   
+         else let d = z-z' in     
+              if neg d then 0 else nat d)"
+apply (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0)
+done
+
+lemma diff_nat_number_of [simp]: 
+     "(number_of v :: nat) - number_of v' =  
+        (if neg (number_of v' :: int) then number_of v  
+         else let d = number_of (v + uminus v') in     
+              if neg d then 0 else nat d)"
+by (simp del: nat_number_of add: diff_nat_eq_if nat_number_of_def) 
+
+
+
+subsubsection{*Multiplication *}
+
+lemma mult_nat_number_of [simp]:
+     "(number_of v :: nat) * number_of v' =  
+       (if neg (number_of v :: int) then 0 else number_of (v * v'))"
+by (force dest!: neg_nat
+          simp del: nat_number_of
+          simp add: nat_number_of_def nat_mult_distrib [symmetric]) 
+
+
+
+subsubsection{*Quotient *}
+
+lemma div_nat_number_of [simp]:
+     "(number_of v :: nat)  div  number_of v' =  
+          (if neg (number_of v :: int) then 0  
+           else nat (number_of v div number_of v'))"
+by (force dest!: neg_nat
+          simp del: nat_number_of
+          simp add: nat_number_of_def nat_div_distrib [symmetric]) 
+
+lemma one_div_nat_number_of [simp]:
+     "(Suc 0)  div  number_of v' = (nat (1 div number_of v'))" 
+by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
+
+
+subsubsection{*Remainder *}
+
+lemma mod_nat_number_of [simp]:
+     "(number_of v :: nat)  mod  number_of v' =  
+        (if neg (number_of v :: int) then 0  
+         else if neg (number_of v' :: int) then number_of v  
+         else nat (number_of v mod number_of v'))"
+by (force dest!: neg_nat
+          simp del: nat_number_of
+          simp add: nat_number_of_def nat_mod_distrib [symmetric]) 
+
+lemma one_mod_nat_number_of [simp]:
+     "(Suc 0)  mod  number_of v' =  
+        (if neg (number_of v' :: int) then Suc 0
+         else nat (1 mod number_of v'))"
+by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) 
+
+
+subsubsection{* Divisibility *}
+
+lemmas dvd_eq_mod_eq_0_number_of =
+  dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard]
+
+declare dvd_eq_mod_eq_0_number_of [simp]
+
+ML
+{*
+val nat_number_of_def = thm"nat_number_of_def";
+
+val nat_number_of = thm"nat_number_of";
+val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0";
+val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1";
+val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0";
+val numeral_2_eq_2 = thm"numeral_2_eq_2";
+val nat_div_distrib = thm"nat_div_distrib";
+val nat_mod_distrib = thm"nat_mod_distrib";
+val int_nat_number_of = thm"int_nat_number_of";
+val Suc_nat_eq_nat_zadd1 = thm"Suc_nat_eq_nat_zadd1";
+val Suc_nat_number_of_add = thm"Suc_nat_number_of_add";
+val Suc_nat_number_of = thm"Suc_nat_number_of";
+val add_nat_number_of = thm"add_nat_number_of";
+val diff_nat_eq_if = thm"diff_nat_eq_if";
+val diff_nat_number_of = thm"diff_nat_number_of";
+val mult_nat_number_of = thm"mult_nat_number_of";
+val div_nat_number_of = thm"div_nat_number_of";
+val mod_nat_number_of = thm"mod_nat_number_of";
+*}
+
+
+subsection{*Comparisons*}
+
+subsubsection{*Equals (=) *}
+
+lemma eq_nat_nat_iff:
+     "[| (0::int) <= z;  0 <= z' |] ==> (nat z = nat z') = (z=z')"
+by (auto elim!: nonneg_eq_int)
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+lemma eq_nat_number_of [simp]:
+     "((number_of v :: nat) = number_of v') =  
+      (if neg (number_of v :: int) then (iszero (number_of v' :: int) | neg (number_of v' :: int))  
+       else if neg (number_of v' :: int) then iszero (number_of v :: int)  
+       else iszero (number_of (v + uminus v') :: int))"
+apply (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
+                  eq_nat_nat_iff eq_number_of_eq nat_0 iszero_def
+            split add: split_if cong add: imp_cong)
+apply (simp only: nat_eq_iff nat_eq_iff2)
+apply (simp add: not_neg_eq_ge_0 [symmetric])
+done
+
+
+subsubsection{*Less-than (<) *}
+
+(*"neg" is used in rewrite rules for binary comparisons*)
+lemma less_nat_number_of [simp]:
+     "((number_of v :: nat) < number_of v') =  
+         (if neg (number_of v :: int) then neg (number_of (uminus v') :: int)  
+          else neg (number_of (v + uminus v') :: int))"
+by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def
+                nat_less_eq_zless less_number_of_eq_neg zless_nat_eq_int_zless
+         cong add: imp_cong, simp add: Pls_def)
+
+
+(*Maps #n to n for n = 0, 1, 2*)
+lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2
+
+
+subsection{*Powers with Numeric Exponents*}
+
+text{*We cannot refer to the number @{term 2} in @{text Ring_and_Field.thy}.
+We cannot prove general results about the numeral @{term "-1"}, so we have to
+use @{term "- 1"} instead.*}
+
+lemma power2_eq_square: "(a::'a::{comm_semiring_1_cancel,recpower})\<twosuperior> = a * a"
+  by (simp add: numeral_2_eq_2 Power.power_Suc)
+
+lemma zero_power2 [simp]: "(0::'a::{comm_semiring_1_cancel,recpower})\<twosuperior> = 0"
+  by (simp add: power2_eq_square)
+
+lemma one_power2 [simp]: "(1::'a::{comm_semiring_1_cancel,recpower})\<twosuperior> = 1"
+  by (simp add: power2_eq_square)
+
+lemma power3_eq_cube: "(x::'a::recpower) ^ 3 = x * x * x"
+  apply (subgoal_tac "3 = Suc (Suc (Suc 0))")
+  apply (erule ssubst)
+  apply (simp add: power_Suc mult_ac)
+  apply (unfold nat_number_of_def)
+  apply (subst nat_eq_iff)
+  apply simp
+done
+
+text{*Squares of literal numerals will be evaluated.*}
+lemmas power2_eq_square_number_of =
+    power2_eq_square [of "number_of w", standard]
+declare power2_eq_square_number_of [simp]
+
+
+lemma zero_le_power2[simp]: "0 \<le> (a\<twosuperior>::'a::{ordered_idom,recpower})"
+  by (simp add: power2_eq_square)
+
+lemma zero_less_power2[simp]:
+     "(0 < a\<twosuperior>) = (a \<noteq> (0::'a::{ordered_idom,recpower}))"
+  by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
+
+lemma power2_less_0[simp]:
+  fixes a :: "'a::{ordered_idom,recpower}"
+  shows "~ (a\<twosuperior> < 0)"
+by (force simp add: power2_eq_square mult_less_0_iff) 
+
+lemma zero_eq_power2[simp]:
+     "(a\<twosuperior> = 0) = (a = (0::'a::{ordered_idom,recpower}))"
+  by (force simp add: power2_eq_square mult_eq_0_iff)
+
+lemma abs_power2[simp]:
+     "abs(a\<twosuperior>) = (a\<twosuperior>::'a::{ordered_idom,recpower})"
+  by (simp add: power2_eq_square abs_mult abs_mult_self)
+
+lemma power2_abs[simp]:
+     "(abs a)\<twosuperior> = (a\<twosuperior>::'a::{ordered_idom,recpower})"
+  by (simp add: power2_eq_square abs_mult_self)
+
+lemma power2_minus[simp]:
+     "(- a)\<twosuperior> = (a\<twosuperior>::'a::{comm_ring_1,recpower})"
+  by (simp add: power2_eq_square)
+
+lemma power2_le_imp_le:
+  fixes x y :: "'a::{ordered_semidom,recpower}"
+  shows "\<lbrakk>x\<twosuperior> \<le> y\<twosuperior>; 0 \<le> y\<rbrakk> \<Longrightarrow> x \<le> y"
+unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)
+
+lemma power2_less_imp_less:
+  fixes x y :: "'a::{ordered_semidom,recpower}"
+  shows "\<lbrakk>x\<twosuperior> < y\<twosuperior>; 0 \<le> y\<rbrakk> \<Longrightarrow> x < y"
+by (rule power_less_imp_less_base)
+
+lemma power2_eq_imp_eq:
+  fixes x y :: "'a::{ordered_semidom,recpower}"
+  shows "\<lbrakk>x\<twosuperior> = y\<twosuperior>; 0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> x = y"
+unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base, simp)
+
+lemma power_minus1_even[simp]: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,recpower})"
+apply (induct "n")
+apply (auto simp add: power_Suc power_add)
+done
+
+lemma power_even_eq: "(a::'a::recpower) ^ (2*n) = (a^n)^2"
+by (subst mult_commute) (simp add: power_mult)
+
+lemma power_odd_eq: "(a::int) ^ Suc(2*n) = a * (a^n)^2"
+by (simp add: power_even_eq) 
+
+lemma power_minus_even [simp]:
+     "(-a) ^ (2*n) = (a::'a::{comm_ring_1,recpower}) ^ (2*n)"
+by (simp add: power_minus1_even power_minus [of a]) 
+
+lemma zero_le_even_power'[simp]:
+     "0 \<le> (a::'a::{ordered_idom,recpower}) ^ (2*n)"
+proof (induct "n")
+  case 0
+    show ?case by (simp add: zero_le_one)
+next
+  case (Suc n)
+    have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" 
+      by (simp add: mult_ac power_add power2_eq_square)
+    thus ?case
+      by (simp add: prems zero_le_mult_iff)
+qed
+
+lemma odd_power_less_zero:
+     "(a::'a::{ordered_idom,recpower}) < 0 ==> a ^ Suc(2*n) < 0"
+proof (induct "n")
+  case 0
+    show ?case by (simp add: Power.power_Suc)
+next
+  case (Suc n)
+    have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" 
+      by (simp add: mult_ac power_add power2_eq_square Power.power_Suc)
+    thus ?case
+      by (simp add: prems mult_less_0_iff mult_neg_neg)
+qed
+
+lemma odd_0_le_power_imp_0_le:
+     "0 \<le> a  ^ Suc(2*n) ==> 0 \<le> (a::'a::{ordered_idom,recpower})"
+apply (insert odd_power_less_zero [of a n]) 
+apply (force simp add: linorder_not_less [symmetric]) 
+done
+
+text{*Simprules for comparisons where common factors can be cancelled.*}
+lemmas zero_compare_simps =
+    add_strict_increasing add_strict_increasing2 add_increasing
+    zero_le_mult_iff zero_le_divide_iff 
+    zero_less_mult_iff zero_less_divide_iff 
+    mult_le_0_iff divide_le_0_iff 
+    mult_less_0_iff divide_less_0_iff 
+    zero_le_power2 power2_less_0
+
+subsubsection{*Nat *}
+
+lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
+by (simp add: numerals)
+
+(*Expresses a natural number constant as the Suc of another one.
+  NOT suitable for rewriting because n recurs in the condition.*)
+lemmas expand_Suc = Suc_pred' [of "number_of v", standard]
+
+subsubsection{*Arith *}
+
+lemma Suc_eq_add_numeral_1: "Suc n = n + 1"
+by (simp add: numerals)
+
+lemma Suc_eq_add_numeral_1_left: "Suc n = 1 + n"
+by (simp add: numerals)
+
+(* These two can be useful when m = number_of... *)
+
+lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))"
+apply (case_tac "m")
+apply (simp_all add: numerals)
+done
+
+lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))"
+apply (case_tac "m")
+apply (simp_all add: numerals)
+done
+
+lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))"
+apply (case_tac "m")
+apply (simp_all add: numerals)
+done
+
+
+subsection{*Comparisons involving (0::nat) *}
+
+text{*Simplification already does @{term "n<0"}, @{term "n\<le>0"} and @{term "0\<le>n"}.*}
+
+lemma eq_number_of_0 [simp]:
+     "(number_of v = (0::nat)) =  
+      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
+by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
+
+lemma eq_0_number_of [simp]:
+     "((0::nat) = number_of v) =  
+      (if neg (number_of v :: int) then True else iszero (number_of v :: int))"
+by (rule trans [OF eq_sym_conv eq_number_of_0])
+
+lemma less_0_number_of [simp]:
+     "((0::nat) < number_of v) = neg (number_of (uminus v) :: int)"
+by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] Pls_def)
+
+
+lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)"
+by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric] iszero_0)
+
+
+
+subsection{*Comparisons involving  @{term Suc} *}
+
+lemma eq_number_of_Suc [simp]:
+     "(number_of v = Suc n) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then False else nat pv = n)"
+apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
+                  number_of_pred nat_number_of_def 
+            split add: split_if)
+apply (rule_tac x = "number_of v" in spec)
+apply (auto simp add: nat_eq_iff)
+done
+
+lemma Suc_eq_number_of [simp]:
+     "(Suc n = number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then False else nat pv = n)"
+by (rule trans [OF eq_sym_conv eq_number_of_Suc])
+
+lemma less_number_of_Suc [simp]:
+     "(number_of v < Suc n) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then True else nat pv < n)"
+apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
+                  number_of_pred nat_number_of_def  
+            split add: split_if)
+apply (rule_tac x = "number_of v" in spec)
+apply (auto simp add: nat_less_iff)
+done
+
+lemma less_Suc_number_of [simp]:
+     "(Suc n < number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then False else n < nat pv)"
+apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less 
+                  number_of_pred nat_number_of_def
+            split add: split_if)
+apply (rule_tac x = "number_of v" in spec)
+apply (auto simp add: zless_nat_eq_int_zless)
+done
+
+lemma le_number_of_Suc [simp]:
+     "(number_of v <= Suc n) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then True else nat pv <= n)"
+by (simp add: Let_def less_Suc_number_of linorder_not_less [symmetric])
+
+lemma le_Suc_number_of [simp]:
+     "(Suc n <= number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then False else n <= nat pv)"
+by (simp add: Let_def less_number_of_Suc linorder_not_less [symmetric])
+
+
+(* Push int(.) inwards: *)
+declare zadd_int [symmetric, simp]
+
+lemma lemma1: "(m+m = n+n) = (m = (n::int))"
+by auto
+
+lemma lemma2: "m+m ~= (1::int) + (n + n)"
+apply auto
+apply (drule_tac f = "%x. x mod 2" in arg_cong)
+apply (simp add: zmod_zadd1_eq)
+done
+
+lemma eq_number_of_BIT_BIT:
+     "((number_of (v BIT x) ::int) = number_of (w BIT y)) =  
+      (x=y & (((number_of v) ::int) = number_of w))"
+apply (simp only: number_of_BIT lemma1 lemma2 eq_commute
+               OrderedGroup.add_left_cancel add_assoc OrderedGroup.add_0_left
+            split add: bit.split)
+apply simp
+done
+
+lemma eq_number_of_BIT_Pls:
+     "((number_of (v BIT x) ::int) = Numeral0) =  
+      (x=bit.B0 & (((number_of v) ::int) = Numeral0))"
+apply (simp only: simp_thms  add: number_of_BIT number_of_Pls eq_commute
+            split add: bit.split cong: imp_cong)
+apply (rule_tac x = "number_of v" in spec, safe)
+apply (simp_all (no_asm_use))
+apply (drule_tac f = "%x. x mod 2" in arg_cong)
+apply (simp add: zmod_zadd1_eq)
+done
+
+lemma eq_number_of_BIT_Min:
+     "((number_of (v BIT x) ::int) = number_of Numeral.Min) =  
+      (x=bit.B1 & (((number_of v) ::int) = number_of Numeral.Min))"
+apply (simp only: simp_thms  add: number_of_BIT number_of_Min eq_commute
+            split add: bit.split cong: imp_cong)
+apply (rule_tac x = "number_of v" in spec, auto)
+apply (drule_tac f = "%x. x mod 2" in arg_cong, auto)
+done
+
+lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Numeral.Min"
+by auto
+
+
+
+subsection{*Max and Min Combined with @{term Suc} *}
+
+lemma max_number_of_Suc [simp]:
+     "max (Suc n) (number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then Suc n else Suc(max n (nat pv)))"
+apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
+            split add: split_if nat.split)
+apply (rule_tac x = "number_of v" in spec) 
+apply auto
+done
+ 
+lemma max_Suc_number_of [simp]:
+     "max (number_of v) (Suc n) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then Suc n else Suc(max (nat pv) n))"
+apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
+            split add: split_if nat.split)
+apply (rule_tac x = "number_of v" in spec) 
+apply auto
+done
+ 
+lemma min_number_of_Suc [simp]:
+     "min (Suc n) (number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then 0 else Suc(min n (nat pv)))"
+apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
+            split add: split_if nat.split)
+apply (rule_tac x = "number_of v" in spec) 
+apply auto
+done
+ 
+lemma min_Suc_number_of [simp]:
+     "min (number_of v) (Suc n) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then 0 else Suc(min (nat pv) n))"
+apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def 
+            split add: split_if nat.split)
+apply (rule_tac x = "number_of v" in spec) 
+apply auto
+done
+ 
+subsection{*Literal arithmetic involving powers*}
+
+lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n"
+apply (induct "n")
+apply (simp_all (no_asm_simp) add: nat_mult_distrib)
+done
+
+lemma power_nat_number_of:
+     "(number_of v :: nat) ^ n =  
+       (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))"
+by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq
+         split add: split_if cong: imp_cong)
+
+
+lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w", standard]
+declare power_nat_number_of_number_of [simp]
+
+
+
+text{*For the integers*}
+
+lemma zpower_number_of_even:
+  "(z::int) ^ number_of (w BIT bit.B0) = (let w = z ^ (number_of w) in w * w)"
+unfolding Let_def nat_number_of_def number_of_BIT bit.cases
+apply (rule_tac x = "number_of w" in spec, clarify)
+apply (case_tac " (0::int) <= x")
+apply (auto simp add: nat_mult_distrib power_even_eq power2_eq_square)
+done
+
+lemma zpower_number_of_odd:
+  "(z::int) ^ number_of (w BIT bit.B1) = (if (0::int) <= number_of w                    
+     then (let w = z ^ (number_of w) in z * w * w) else 1)"
+unfolding Let_def nat_number_of_def number_of_BIT bit.cases
+apply (rule_tac x = "number_of w" in spec, auto)
+apply (simp only: nat_add_distrib nat_mult_distrib)
+apply simp
+apply (auto simp add: nat_add_distrib nat_mult_distrib power_even_eq power2_eq_square neg_nat)
+done
+
+lemmas zpower_number_of_even_number_of =
+    zpower_number_of_even [of "number_of v", standard]
+declare zpower_number_of_even_number_of [simp]
+
+lemmas zpower_number_of_odd_number_of =
+    zpower_number_of_odd [of "number_of v", standard]
+declare zpower_number_of_odd_number_of [simp]
+
+
+
+
+ML
+{*
+val numerals = thms"numerals";
+val numeral_ss = simpset() addsimps numerals;
+
+val nat_bin_arith_setup =
+ Fast_Arith.map_data
+   (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+     {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
+      inj_thms = inj_thms,
+      lessD = lessD, neqE = neqE,
+      simpset = simpset addsimps [Suc_nat_number_of, int_nat_number_of,
+                                  not_neg_number_of_Pls,
+                                  neg_number_of_Min,neg_number_of_BIT]})
+*}
+
+setup nat_bin_arith_setup
+
+(* Enable arith to deal with div/mod k where k is a numeral: *)
+declare split_div[of _ _ "number_of k", standard, arith_split]
+declare split_mod[of _ _ "number_of k", standard, arith_split]
+
+lemma nat_number_of_Pls: "Numeral0 = (0::nat)"
+  by (simp add: number_of_Pls nat_number_of_def)
+
+lemma nat_number_of_Min: "number_of Numeral.Min = (0::nat)"
+  apply (simp only: number_of_Min nat_number_of_def nat_zminus_int)
+  done
+
+lemma nat_number_of_BIT_1:
+  "number_of (w BIT bit.B1) =
+    (if neg (number_of w :: int) then 0
+     else let n = number_of w in Suc (n + n))"
+  apply (simp only: nat_number_of_def Let_def split: split_if)
+  apply (intro conjI impI)
+   apply (simp add: neg_nat neg_number_of_BIT)
+  apply (rule int_int_eq [THEN iffD1])
+  apply (simp only: not_neg_nat neg_number_of_BIT int_Suc zadd_int [symmetric] simp_thms)
+  apply (simp only: number_of_BIT zadd_assoc split: bit.split)
+  apply simp
+  done
+
+lemma nat_number_of_BIT_0:
+    "number_of (w BIT bit.B0) = (let n::nat = number_of w in n + n)"
+  apply (simp only: nat_number_of_def Let_def)
+  apply (cases "neg (number_of w :: int)")
+   apply (simp add: neg_nat neg_number_of_BIT)
+  apply (rule int_int_eq [THEN iffD1])
+  apply (simp only: not_neg_nat neg_number_of_BIT int_Suc zadd_int [symmetric] simp_thms)
+  apply (simp only: number_of_BIT zadd_assoc)
+  apply simp
+  done
+
+lemmas nat_number =
+  nat_number_of_Pls nat_number_of_Min
+  nat_number_of_BIT_1 nat_number_of_BIT_0
+
+lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
+  by (simp add: Let_def)
+
+lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring,recpower})"
+by (simp add: power_mult); 
+
+lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring,recpower})"
+by (simp add: power_mult power_Suc); 
+
+
+subsection{*Literal arithmetic and @{term of_nat}*}
+
+lemma of_nat_double:
+     "0 \<le> x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)"
+by (simp only: mult_2 nat_add_distrib of_nat_add) 
+
+lemma nat_numeral_m1_eq_0: "-1 = (0::nat)"
+by (simp only: nat_number_of_def)
+
+lemma of_nat_number_of_lemma:
+     "of_nat (number_of v :: nat) =  
+         (if 0 \<le> (number_of v :: int) 
+          then (number_of v :: 'a :: number_ring)
+          else 0)"
+by (simp add: int_number_of_def nat_number_of_def number_of_eq of_nat_nat);
+
+lemma of_nat_number_of_eq [simp]:
+     "of_nat (number_of v :: nat) =  
+         (if neg (number_of v :: int) then 0  
+          else (number_of v :: 'a :: number_ring))"
+by (simp only: of_nat_number_of_lemma neg_def, simp) 
+
+
+subsection {*Lemmas for the Combination and Cancellation Simprocs*}
+
+lemma nat_number_of_add_left:
+     "number_of v + (number_of v' + (k::nat)) =  
+         (if neg (number_of v :: int) then number_of v' + k  
+          else if neg (number_of v' :: int) then number_of v + k  
+          else number_of (v + v') + k)"
+by simp
+
+lemma nat_number_of_mult_left:
+     "number_of v * (number_of v' * (k::nat)) =  
+         (if neg (number_of v :: int) then 0
+          else number_of (v * v') * k)"
+by simp
+
+
+subsubsection{*For @{text combine_numerals}*}
+
+lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)"
+by (simp add: add_mult_distrib)
+
+
+subsubsection{*For @{text cancel_numerals}*}
+
+lemma nat_diff_add_eq1:
+     "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)"
+by (simp split add: nat_diff_split add: add_mult_distrib)
+
+lemma nat_diff_add_eq2:
+     "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))"
+by (simp split add: nat_diff_split add: add_mult_distrib)
+
+lemma nat_eq_add_iff1:
+     "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)"
+by (auto split add: nat_diff_split simp add: add_mult_distrib)
+
+lemma nat_eq_add_iff2:
+     "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)"
+by (auto split add: nat_diff_split simp add: add_mult_distrib)
+
+lemma nat_less_add_iff1:
+     "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)"
+by (auto split add: nat_diff_split simp add: add_mult_distrib)
+
+lemma nat_less_add_iff2:
+     "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)"
+by (auto split add: nat_diff_split simp add: add_mult_distrib)
+
+lemma nat_le_add_iff1:
+     "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)"
+by (auto split add: nat_diff_split simp add: add_mult_distrib)
+
+lemma nat_le_add_iff2:
+     "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)"
+by (auto split add: nat_diff_split simp add: add_mult_distrib)
+
+
+subsubsection{*For @{text cancel_numeral_factors} *}
+
+lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)"
+by auto
+
+lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m<n)"
+by auto
+
+lemma nat_mult_eq_cancel1: "(0::nat) < k ==> (k*m = k*n) = (m=n)"
+by auto
+
+lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)"
+by auto
+
+
+subsubsection{*For @{text cancel_factor} *}
+
+lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)"
+by auto
+
+lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m<n)"
+by auto
+
+lemma nat_mult_eq_cancel_disj: "(k*m = k*n) = (k = (0::nat) | m=n)"
+by auto
+
+lemma nat_mult_div_cancel_disj:
+     "(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
+by (simp add: nat_mult_div_cancel1)
+
+
+subsection {* legacy ML bindings *}
+
+ML
+{*
+val eq_nat_nat_iff = thm"eq_nat_nat_iff";
+val eq_nat_number_of = thm"eq_nat_number_of";
+val less_nat_number_of = thm"less_nat_number_of";
+val power2_eq_square = thm "power2_eq_square";
+val zero_le_power2 = thm "zero_le_power2";
+val zero_less_power2 = thm "zero_less_power2";
+val zero_eq_power2 = thm "zero_eq_power2";
+val abs_power2 = thm "abs_power2";
+val power2_abs = thm "power2_abs";
+val power2_minus = thm "power2_minus";
+val power_minus1_even = thm "power_minus1_even";
+val power_minus_even = thm "power_minus_even";
+val odd_power_less_zero = thm "odd_power_less_zero";
+val odd_0_le_power_imp_0_le = thm "odd_0_le_power_imp_0_le";
+
+val Suc_pred' = thm"Suc_pred'";
+val expand_Suc = thm"expand_Suc";
+val Suc_eq_add_numeral_1 = thm"Suc_eq_add_numeral_1";
+val Suc_eq_add_numeral_1_left = thm"Suc_eq_add_numeral_1_left";
+val add_eq_if = thm"add_eq_if";
+val mult_eq_if = thm"mult_eq_if";
+val power_eq_if = thm"power_eq_if";
+val eq_number_of_0 = thm"eq_number_of_0";
+val eq_0_number_of = thm"eq_0_number_of";
+val less_0_number_of = thm"less_0_number_of";
+val neg_imp_number_of_eq_0 = thm"neg_imp_number_of_eq_0";
+val eq_number_of_Suc = thm"eq_number_of_Suc";
+val Suc_eq_number_of = thm"Suc_eq_number_of";
+val less_number_of_Suc = thm"less_number_of_Suc";
+val less_Suc_number_of = thm"less_Suc_number_of";
+val le_number_of_Suc = thm"le_number_of_Suc";
+val le_Suc_number_of = thm"le_Suc_number_of";
+val eq_number_of_BIT_BIT = thm"eq_number_of_BIT_BIT";
+val eq_number_of_BIT_Pls = thm"eq_number_of_BIT_Pls";
+val eq_number_of_BIT_Min = thm"eq_number_of_BIT_Min";
+val eq_number_of_Pls_Min = thm"eq_number_of_Pls_Min";
+val of_nat_number_of_eq = thm"of_nat_number_of_eq";
+val nat_power_eq = thm"nat_power_eq";
+val power_nat_number_of = thm"power_nat_number_of";
+val zpower_number_of_even = thm"zpower_number_of_even";
+val zpower_number_of_odd = thm"zpower_number_of_odd";
+val nat_number_of_Pls = thm"nat_number_of_Pls";
+val nat_number_of_Min = thm"nat_number_of_Min";
+val Let_Suc = thm"Let_Suc";
+
+val nat_number = thms"nat_number";
+
+val nat_number_of_add_left = thm"nat_number_of_add_left";
+val nat_number_of_mult_left = thm"nat_number_of_mult_left";
+val left_add_mult_distrib = thm"left_add_mult_distrib";
+val nat_diff_add_eq1 = thm"nat_diff_add_eq1";
+val nat_diff_add_eq2 = thm"nat_diff_add_eq2";
+val nat_eq_add_iff1 = thm"nat_eq_add_iff1";
+val nat_eq_add_iff2 = thm"nat_eq_add_iff2";
+val nat_less_add_iff1 = thm"nat_less_add_iff1";
+val nat_less_add_iff2 = thm"nat_less_add_iff2";
+val nat_le_add_iff1 = thm"nat_le_add_iff1";
+val nat_le_add_iff2 = thm"nat_le_add_iff2";
+val nat_mult_le_cancel1 = thm"nat_mult_le_cancel1";
+val nat_mult_less_cancel1 = thm"nat_mult_less_cancel1";
+val nat_mult_eq_cancel1 = thm"nat_mult_eq_cancel1";
+val nat_mult_div_cancel1 = thm"nat_mult_div_cancel1";
+val nat_mult_le_cancel_disj = thm"nat_mult_le_cancel_disj";
+val nat_mult_less_cancel_disj = thm"nat_mult_less_cancel_disj";
+val nat_mult_eq_cancel_disj = thm"nat_mult_eq_cancel_disj";
+val nat_mult_div_cancel_disj = thm"nat_mult_div_cancel_disj";
+
+val power_minus_even = thm"power_minus_even";
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/NatSimprocs.thy	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,391 @@
+(*  Title:      HOL/NatSimprocs.thy
+    ID:         $Id$
+    Copyright   2003 TU Muenchen
+*)
+
+header {*Simprocs for the Naturals*}
+
+theory NatSimprocs
+imports NatBin
+uses "int_factor_simprocs.ML" "nat_simprocs.ML"
+begin
+
+setup nat_simprocs_setup
+
+subsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
+
+text{*Where K above is a literal*}
+
+lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
+by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
+
+text {*Now just instantiating @{text n} to @{text "number_of v"} does
+  the right simplification, but with some redundant inequality
+  tests.*}
+lemma neg_number_of_pred_iff_0:
+  "neg (number_of (Numeral.pred v)::int) = (number_of v = (0::nat))"
+apply (subgoal_tac "neg (number_of (Numeral.pred v)) = (number_of v < Suc 0) ")
+apply (simp only: less_Suc_eq_le le_0_eq)
+apply (subst less_number_of_Suc, simp)
+done
+
+text{*No longer required as a simprule because of the @{text inverse_fold}
+   simproc*}
+lemma Suc_diff_number_of:
+     "neg (number_of (uminus v)::int) ==>  
+      Suc m - (number_of v) = m - (number_of (Numeral.pred v))"
+apply (subst Suc_diff_eq_diff_pred)
+apply simp
+apply (simp del: nat_numeral_1_eq_1)
+apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] 
+                        neg_number_of_pred_iff_0)
+done
+
+lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
+by (simp add: numerals split add: nat_diff_split)
+
+
+subsection{*For @{term nat_case} and @{term nat_rec}*}
+
+lemma nat_case_number_of [simp]:
+     "nat_case a f (number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then a else f (nat pv))"
+by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
+
+lemma nat_case_add_eq_if [simp]:
+     "nat_case a f ((number_of v) + n) =  
+       (let pv = number_of (Numeral.pred v) in  
+         if neg pv then nat_case a f n else f (nat pv + n))"
+apply (subst add_eq_if)
+apply (simp split add: nat.split
+            del: nat_numeral_1_eq_1
+	    add: numeral_1_eq_Suc_0 [symmetric] Let_def 
+                 neg_imp_number_of_eq_0 neg_number_of_pred_iff_0)
+done
+
+lemma nat_rec_number_of [simp]:
+     "nat_rec a f (number_of v) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
+apply (case_tac " (number_of v) ::nat")
+apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
+apply (simp split add: split_if_asm)
+done
+
+lemma nat_rec_add_eq_if [simp]:
+     "nat_rec a f (number_of v + n) =  
+        (let pv = number_of (Numeral.pred v) in  
+         if neg pv then nat_rec a f n  
+                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
+apply (subst add_eq_if)
+apply (simp split add: nat.split
+            del: nat_numeral_1_eq_1
+            add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
+                 neg_number_of_pred_iff_0)
+done
+
+
+subsection{*Various Other Lemmas*}
+
+subsubsection{*Evens and Odds, for Mutilated Chess Board*}
+
+text{*Lemmas for specialist use, NOT as default simprules*}
+lemma nat_mult_2: "2 * z = (z+z::nat)"
+proof -
+  have "2*z = (1 + 1)*z" by simp
+  also have "... = z+z" by (simp add: left_distrib)
+  finally show ?thesis .
+qed
+
+lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
+by (subst mult_commute, rule nat_mult_2)
+
+text{*Case analysis on @{term "n<2"}*}
+lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
+by arith
+
+lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
+by arith
+
+lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
+by (simp add: nat_mult_2 [symmetric])
+
+lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
+apply (subgoal_tac "m mod 2 < 2")
+apply (erule less_2_cases [THEN disjE])
+apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
+done
+
+lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
+apply (subgoal_tac "m mod 2 < 2")
+apply (force simp del: mod_less_divisor, simp) 
+done
+
+subsubsection{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
+
+lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
+by simp
+
+lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
+by simp
+
+text{*Can be used to eliminate long strings of Sucs, but not by default*}
+lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
+by simp
+
+
+text{*These lemmas collapse some needless occurrences of Suc:
+    at least three Sucs, since two and fewer are rewritten back to Suc again!
+    We already have some rules to simplify operands smaller than 3.*}
+
+lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
+by (simp add: Suc3_eq_add_3)
+
+lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
+by (simp add: Suc3_eq_add_3)
+
+lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
+by (simp add: Suc3_eq_add_3)
+
+lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
+by (simp add: Suc3_eq_add_3)
+
+lemmas Suc_div_eq_add3_div_number_of =
+    Suc_div_eq_add3_div [of _ "number_of v", standard]
+declare Suc_div_eq_add3_div_number_of [simp]
+
+lemmas Suc_mod_eq_add3_mod_number_of =
+    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
+declare Suc_mod_eq_add3_mod_number_of [simp]
+
+
+
+subsection{*Special Simplification for Constants*}
+
+text{*These belong here, late in the development of HOL, to prevent their
+interfering with proofs of abstract properties of instances of the function
+@{term number_of}*}
+
+text{*These distributive laws move literals inside sums and differences.*}
+lemmas left_distrib_number_of = left_distrib [of _ _ "number_of v", standard]
+declare left_distrib_number_of [simp]
+
+lemmas right_distrib_number_of = right_distrib [of "number_of v", standard]
+declare right_distrib_number_of [simp]
+
+
+lemmas left_diff_distrib_number_of =
+    left_diff_distrib [of _ _ "number_of v", standard]
+declare left_diff_distrib_number_of [simp]
+
+lemmas right_diff_distrib_number_of =
+    right_diff_distrib [of "number_of v", standard]
+declare right_diff_distrib_number_of [simp]
+
+
+text{*These are actually for fields, like real: but where else to put them?*}
+lemmas zero_less_divide_iff_number_of =
+    zero_less_divide_iff [of "number_of w", standard]
+declare zero_less_divide_iff_number_of [simp]
+
+lemmas divide_less_0_iff_number_of =
+    divide_less_0_iff [of "number_of w", standard]
+declare divide_less_0_iff_number_of [simp]
+
+lemmas zero_le_divide_iff_number_of =
+    zero_le_divide_iff [of "number_of w", standard]
+declare zero_le_divide_iff_number_of [simp]
+
+lemmas divide_le_0_iff_number_of =
+    divide_le_0_iff [of "number_of w", standard]
+declare divide_le_0_iff_number_of [simp]
+
+
+(****
+IF times_divide_eq_right and times_divide_eq_left are removed as simprules,
+then these special-case declarations may be useful.
+
+text{*These simprules move numerals into numerators and denominators.*}
+lemma times_recip_eq_right [simp]: "a * (1/c) = a / (c::'a::field)"
+by (simp add: times_divide_eq)
+
+lemma times_recip_eq_left [simp]: "(1/c) * a = a / (c::'a::field)"
+by (simp add: times_divide_eq)
+
+lemmas times_divide_eq_right_number_of =
+    times_divide_eq_right [of "number_of w", standard]
+declare times_divide_eq_right_number_of [simp]
+
+lemmas times_divide_eq_right_number_of =
+    times_divide_eq_right [of _ _ "number_of w", standard]
+declare times_divide_eq_right_number_of [simp]
+
+lemmas times_divide_eq_left_number_of =
+    times_divide_eq_left [of _ "number_of w", standard]
+declare times_divide_eq_left_number_of [simp]
+
+lemmas times_divide_eq_left_number_of =
+    times_divide_eq_left [of _ _ "number_of w", standard]
+declare times_divide_eq_left_number_of [simp]
+
+****)
+
+text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
+  strange, but then other simprocs simplify the quotient.*}
+
+lemmas inverse_eq_divide_number_of =
+    inverse_eq_divide [of "number_of w", standard]
+declare inverse_eq_divide_number_of [simp]
+
+
+subsubsection{*These laws simplify inequalities, moving unary minus from a term
+into the literal.*}
+lemmas less_minus_iff_number_of =
+    less_minus_iff [of "number_of v", standard]
+declare less_minus_iff_number_of [simp]
+
+lemmas le_minus_iff_number_of =
+    le_minus_iff [of "number_of v", standard]
+declare le_minus_iff_number_of [simp]
+
+lemmas equation_minus_iff_number_of =
+    equation_minus_iff [of "number_of v", standard]
+declare equation_minus_iff_number_of [simp]
+
+
+lemmas minus_less_iff_number_of =
+    minus_less_iff [of _ "number_of v", standard]
+declare minus_less_iff_number_of [simp]
+
+lemmas minus_le_iff_number_of =
+    minus_le_iff [of _ "number_of v", standard]
+declare minus_le_iff_number_of [simp]
+
+lemmas minus_equation_iff_number_of =
+    minus_equation_iff [of _ "number_of v", standard]
+declare minus_equation_iff_number_of [simp]
+
+
+subsubsection{*To Simplify Inequalities Where One Side is the Constant 1*}
+
+lemma less_minus_iff_1 [simp]: 
+  fixes b::"'b::{ordered_idom,number_ring}" 
+  shows "(1 < - b) = (b < -1)"
+by auto
+
+lemma le_minus_iff_1 [simp]: 
+  fixes b::"'b::{ordered_idom,number_ring}" 
+  shows "(1 \<le> - b) = (b \<le> -1)"
+by auto
+
+lemma equation_minus_iff_1 [simp]: 
+  fixes b::"'b::number_ring" 
+  shows "(1 = - b) = (b = -1)"
+by (subst equation_minus_iff, auto) 
+
+lemma minus_less_iff_1 [simp]: 
+  fixes a::"'b::{ordered_idom,number_ring}" 
+  shows "(- a < 1) = (-1 < a)"
+by auto
+
+lemma minus_le_iff_1 [simp]: 
+  fixes a::"'b::{ordered_idom,number_ring}" 
+  shows "(- a \<le> 1) = (-1 \<le> a)"
+by auto
+
+lemma minus_equation_iff_1 [simp]: 
+  fixes a::"'b::number_ring" 
+  shows "(- a = 1) = (a = -1)"
+by (subst minus_equation_iff, auto) 
+
+
+subsubsection {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
+
+lemmas mult_less_cancel_left_number_of =
+    mult_less_cancel_left [of "number_of v", standard]
+declare mult_less_cancel_left_number_of [simp]
+
+lemmas mult_less_cancel_right_number_of =
+    mult_less_cancel_right [of _ "number_of v", standard]
+declare mult_less_cancel_right_number_of [simp]
+
+lemmas mult_le_cancel_left_number_of =
+    mult_le_cancel_left [of "number_of v", standard]
+declare mult_le_cancel_left_number_of [simp]
+
+lemmas mult_le_cancel_right_number_of =
+    mult_le_cancel_right [of _ "number_of v", standard]
+declare mult_le_cancel_right_number_of [simp]
+
+
+subsubsection {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
+
+lemmas le_divide_eq_number_of = le_divide_eq [of _ _ "number_of w", standard]
+declare le_divide_eq_number_of [simp]
+
+lemmas divide_le_eq_number_of = divide_le_eq [of _ "number_of w", standard]
+declare divide_le_eq_number_of [simp]
+
+lemmas less_divide_eq_number_of = less_divide_eq [of _ _ "number_of w", standard]
+declare less_divide_eq_number_of [simp]
+
+lemmas divide_less_eq_number_of = divide_less_eq [of _ "number_of w", standard]
+declare divide_less_eq_number_of [simp]
+
+lemmas eq_divide_eq_number_of = eq_divide_eq [of _ _ "number_of w", standard]
+declare eq_divide_eq_number_of [simp]
+
+lemmas divide_eq_eq_number_of = divide_eq_eq [of _ "number_of w", standard]
+declare divide_eq_eq_number_of [simp]
+
+
+
+subsection{*Optional Simplification Rules Involving Constants*}
+
+text{*Simplify quotients that are compared with a literal constant.*}
+
+lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard]
+lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard]
+lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard]
+lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard]
+lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard]
+lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard]
+
+
+text{*Not good as automatic simprules because they cause case splits.*}
+lemmas divide_const_simps =
+  le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of
+  divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of
+  le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
+
+subsubsection{*Division By @{text "-1"}*}
+
+lemma divide_minus1 [simp]:
+     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})" 
+by simp
+
+lemma minus1_divide [simp]:
+     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
+by (simp add: divide_inverse inverse_minus_eq)
+
+lemma half_gt_zero_iff:
+     "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
+by auto
+
+lemmas half_gt_zero = half_gt_zero_iff [THEN iffD2, standard]
+declare half_gt_zero [simp]
+
+(* The following lemma should appear in Divides.thy, but there the proof
+   doesn't work. *)
+
+lemma nat_dvd_not_less:
+  "[| 0 < m; m < n |] ==> \<not> n dvd (m::nat)"
+  by (unfold dvd_def) auto
+
+ML {*
+val divide_minus1 = @{thm divide_minus1};
+val minus1_divide = @{thm minus1_divide};
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Numeral.thy	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,685 @@
+(*  Title:      HOL/Numeral.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1994  University of Cambridge
+*)
+
+header {* Arithmetic on Binary Integers *}
+
+theory Numeral
+imports IntDef
+uses ("Tools/numeral_syntax.ML")
+begin
+
+subsection {* Binary representation *}
+
+text {*
+  This formalization defines binary arithmetic in terms of the integers
+  rather than using a datatype. This avoids multiple representations (leading
+  zeroes, etc.)  See @{text "ZF/Tools/twos-compl.ML"}, function @{text
+  int_of_binary}, for the numerical interpretation.
+
+  The representation expects that @{text "(m mod 2)"} is 0 or 1,
+  even if m is negative;
+  For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
+  @{text "-5 = (-3)*2 + 1"}.
+*}
+
+datatype bit = B0 | B1
+
+text{*
+  Type @{typ bit} avoids the use of type @{typ bool}, which would make
+  all of the rewrite rules higher-order.
+*}
+
+definition
+  Pls :: int where
+  [code func del]:"Pls = 0"
+
+definition
+  Min :: int where
+  [code func del]:"Min = - 1"
+
+definition
+  Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where
+  [code func del]: "k BIT b = (case b of B0 \<Rightarrow> 0 | B1 \<Rightarrow> 1) + k + k"
+
+class number = type + -- {* for numeric types: nat, int, real, \dots *}
+  fixes number_of :: "int \<Rightarrow> 'a"
+
+syntax
+  "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
+
+use "Tools/numeral_syntax.ML"
+setup NumeralSyntax.setup
+
+abbreviation
+  "Numeral0 \<equiv> number_of Pls"
+
+abbreviation
+  "Numeral1 \<equiv> number_of (Pls BIT B1)"
+
+lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)"
+  -- {* Unfold all @{text let}s involving constants *}
+  unfolding Let_def ..
+
+lemma Let_0 [simp]: "Let 0 f = f 0"
+  unfolding Let_def ..
+
+lemma Let_1 [simp]: "Let 1 f = f 1"
+  unfolding Let_def ..
+
+definition
+  succ :: "int \<Rightarrow> int" where
+  [code func del]: "succ k = k + 1"
+
+definition
+  pred :: "int \<Rightarrow> int" where
+  [code func del]: "pred k = k - 1"
+
+lemmas
+  max_number_of [simp] = max_def
+    [of "number_of u" "number_of v", standard, simp]
+and
+  min_number_of [simp] = min_def 
+    [of "number_of u" "number_of v", standard, simp]
+  -- {* unfolding @{text minx} and @{text max} on numerals *}
+
+lemmas numeral_simps = 
+  succ_def pred_def Pls_def Min_def Bit_def
+
+text {* Removal of leading zeroes *}
+
+lemma Pls_0_eq [simp, normal post]:
+  "Pls BIT B0 = Pls"
+  unfolding numeral_simps by simp
+
+lemma Min_1_eq [simp, normal post]:
+  "Min BIT B1 = Min"
+  unfolding numeral_simps by simp
+
+
+subsection {* The Functions @{term succ}, @{term pred} and @{term uminus} *}
+
+lemma succ_Pls [simp]:
+  "succ Pls = Pls BIT B1"
+  unfolding numeral_simps by simp
+
+lemma succ_Min [simp]:
+  "succ Min = Pls"
+  unfolding numeral_simps by simp
+
+lemma succ_1 [simp]:
+  "succ (k BIT B1) = succ k BIT B0"
+  unfolding numeral_simps by simp
+
+lemma succ_0 [simp]:
+  "succ (k BIT B0) = k BIT B1"
+  unfolding numeral_simps by simp
+
+lemma pred_Pls [simp]:
+  "pred Pls = Min"
+  unfolding numeral_simps by simp
+
+lemma pred_Min [simp]:
+  "pred Min = Min BIT B0"
+  unfolding numeral_simps by simp
+
+lemma pred_1 [simp]:
+  "pred (k BIT B1) = k BIT B0"
+  unfolding numeral_simps by simp
+
+lemma pred_0 [simp]:
+  "pred (k BIT B0) = pred k BIT B1"
+  unfolding numeral_simps by simp 
+
+lemma minus_Pls [simp]:
+  "- Pls = Pls"
+  unfolding numeral_simps by simp 
+
+lemma minus_Min [simp]:
+  "- Min = Pls BIT B1"
+  unfolding numeral_simps by simp 
+
+lemma minus_1 [simp]:
+  "- (k BIT B1) = pred (- k) BIT B1"
+  unfolding numeral_simps by simp 
+
+lemma minus_0 [simp]:
+  "- (k BIT B0) = (- k) BIT B0"
+  unfolding numeral_simps by simp 
+
+
+subsection {*
+  Binary Addition and Multiplication: @{term "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"}
+    and @{term "op * \<Colon> int \<Rightarrow> int \<Rightarrow> int"}
+*}
+
+lemma add_Pls [simp]:
+  "Pls + k = k"
+  unfolding numeral_simps by simp 
+
+lemma add_Min [simp]:
+  "Min + k = pred k"
+  unfolding numeral_simps by simp
+
+lemma add_BIT_11 [simp]:
+  "(k BIT B1) + (l BIT B1) = (k + succ l) BIT B0"
+  unfolding numeral_simps by simp
+
+lemma add_BIT_10 [simp]:
+  "(k BIT B1) + (l BIT B0) = (k + l) BIT B1"
+  unfolding numeral_simps by simp
+
+lemma add_BIT_0 [simp]:
+  "(k BIT B0) + (l BIT b) = (k + l) BIT b"
+  unfolding numeral_simps by simp 
+
+lemma add_Pls_right [simp]:
+  "k + Pls = k"
+  unfolding numeral_simps by simp 
+
+lemma add_Min_right [simp]:
+  "k + Min = pred k"
+  unfolding numeral_simps by simp 
+
+lemma mult_Pls [simp]:
+  "Pls * w = Pls"
+  unfolding numeral_simps by simp 
+
+lemma mult_Min [simp]:
+  "Min * k = - k"
+  unfolding numeral_simps by simp 
+
+lemma mult_num1 [simp]:
+  "(k BIT B1) * l = ((k * l) BIT B0) + l"
+  unfolding numeral_simps int_distrib by simp 
+
+lemma mult_num0 [simp]:
+  "(k BIT B0) * l = (k * l) BIT B0"
+  unfolding numeral_simps int_distrib by simp 
+
+
+
+subsection {* Converting Numerals to Rings: @{term number_of} *}
+
+axclass number_ring \<subseteq> number, comm_ring_1
+  number_of_eq: "number_of k = of_int k"
+
+text {* self-embedding of the intergers *}
+
+instance int :: number_ring
+  int_number_of_def: "number_of w \<equiv> of_int w"
+  by intro_classes (simp only: int_number_of_def)
+
+lemmas [code func del] = int_number_of_def
+
+lemma number_of_is_id:
+  "number_of (k::int) = k"
+  unfolding int_number_of_def by simp
+
+lemma number_of_succ:
+  "number_of (succ k) = (1 + number_of k ::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma number_of_pred:
+  "number_of (pred w) = (- 1 + number_of w ::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma number_of_minus:
+  "number_of (uminus w) = (- (number_of w)::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma number_of_add:
+  "number_of (v + w) = (number_of v + number_of w::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma number_of_mult:
+  "number_of (v * w) = (number_of v * number_of w::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+text {*
+  The correctness of shifting.
+  But it doesn't seem to give a measurable speed-up.
+*}
+
+lemma double_number_of_BIT:
+  "(1 + 1) * number_of w = (number_of (w BIT B0) ::'a::number_ring)"
+  unfolding number_of_eq numeral_simps left_distrib by simp
+
+text {*
+  Converting numerals 0 and 1 to their abstract versions.
+*}
+
+lemma numeral_0_eq_0 [simp]:
+  "Numeral0 = (0::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma numeral_1_eq_1 [simp]:
+  "Numeral1 = (1::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+text {*
+  Special-case simplification for small constants.
+*}
+
+text{*
+  Unary minus for the abstract constant 1. Cannot be inserted
+  as a simprule until later: it is @{text number_of_Min} re-oriented!
+*}
+
+lemma numeral_m1_eq_minus_1:
+  "(-1::'a::number_ring) = - 1"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma mult_minus1 [simp]:
+  "-1 * z = -(z::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma mult_minus1_right [simp]:
+  "z * -1 = -(z::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+(*Negation of a coefficient*)
+lemma minus_number_of_mult [simp]:
+   "- (number_of w) * z = number_of (uminus w) * (z::'a::number_ring)"
+   unfolding number_of_eq by simp
+
+text {* Subtraction *}
+
+lemma diff_number_of_eq:
+  "number_of v - number_of w =
+    (number_of (v + uminus w)::'a::number_ring)"
+  unfolding number_of_eq by simp
+
+lemma number_of_Pls:
+  "number_of Pls = (0::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma number_of_Min:
+  "number_of Min = (- 1::'a::number_ring)"
+  unfolding number_of_eq numeral_simps by simp
+
+lemma number_of_BIT:
+  "number_of(w BIT x) = (case x of B0 => 0 | B1 => (1::'a::number_ring))
+    + (number_of w) + (number_of w)"
+  unfolding number_of_eq numeral_simps by (simp split: bit.split)
+
+
+subsection {* Equality of Binary Numbers *}
+
+text {* First version by Norbert Voelker *}
+
+lemma eq_number_of_eq:
+  "((number_of x::'a::number_ring) = number_of y) =
+   iszero (number_of (x + uminus y) :: 'a)"
+  unfolding iszero_def number_of_add number_of_minus
+  by (simp add: compare_rls)
+
+lemma iszero_number_of_Pls:
+  "iszero ((number_of Pls)::'a::number_ring)"
+  unfolding iszero_def numeral_0_eq_0 ..
+
+lemma nonzero_number_of_Min:
+  "~ iszero ((number_of Min)::'a::number_ring)"
+  unfolding iszero_def numeral_m1_eq_minus_1 by simp
+
+
+subsection {* Comparisons, for Ordered Rings *}
+
+lemma double_eq_0_iff:
+  "(a + a = 0) = (a = (0::'a::ordered_idom))"
+proof -
+  have "a + a = (1 + 1) * a" unfolding left_distrib by simp
+  with zero_less_two [where 'a = 'a]
+  show ?thesis by force
+qed
+
+lemma le_imp_0_less: 
+  assumes le: "0 \<le> z"
+  shows "(0::int) < 1 + z"
+proof -
+  have "0 \<le> z" .
+  also have "... < z + 1" by (rule less_add_one) 
+  also have "... = 1 + z" by (simp add: add_ac)
+  finally show "0 < 1 + z" .
+qed
+
+lemma odd_nonzero:
+  "1 + z + z \<noteq> (0::int)";
+proof (cases z rule: int_cases)
+  case (nonneg n)
+  have le: "0 \<le> z+z" by (simp add: nonneg add_increasing) 
+  thus ?thesis using  le_imp_0_less [OF le]
+    by (auto simp add: add_assoc) 
+next
+  case (neg n)
+  show ?thesis
+  proof
+    assume eq: "1 + z + z = 0"
+    have "0 < 1 + (int n + int n)"
+      by (simp add: le_imp_0_less add_increasing) 
+    also have "... = - (1 + z + z)" 
+      by (simp add: neg add_assoc [symmetric]) 
+    also have "... = 0" by (simp add: eq) 
+    finally have "0<0" ..
+    thus False by blast
+  qed
+qed
+
+text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
+
+lemma Ints_double_eq_0_iff:
+  assumes in_Ints: "a \<in> Ints"
+  shows "(a + a = 0) = (a = (0::'a::ring_char_0))"
+proof -
+  from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] .
+  then obtain z where a: "a = of_int z" ..
+  show ?thesis
+  proof
+    assume "a = 0"
+    thus "a + a = 0" by simp
+  next
+    assume eq: "a + a = 0"
+    hence "of_int (z + z) = (of_int 0 :: 'a)" by (simp add: a)
+    hence "z + z = 0" by (simp only: of_int_eq_iff)
+    hence "z = 0" by (simp only: double_eq_0_iff)
+    thus "a = 0" by (simp add: a)
+  qed
+qed
+
+lemma Ints_odd_nonzero:
+  assumes in_Ints: "a \<in> Ints"
+  shows "1 + a + a \<noteq> (0::'a::ring_char_0)"
+proof -
+  from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] .
+  then obtain z where a: "a = of_int z" ..
+  show ?thesis
+  proof
+    assume eq: "1 + a + a = 0"
+    hence "of_int (1 + z + z) = (of_int 0 :: 'a)" by (simp add: a)
+    hence "1 + z + z = 0" by (simp only: of_int_eq_iff)
+    with odd_nonzero show False by blast
+  qed
+qed 
+
+lemma Ints_number_of:
+  "(number_of w :: 'a::number_ring) \<in> Ints"
+  unfolding number_of_eq Ints_def by simp
+
+lemma iszero_number_of_BIT:
+  "iszero (number_of (w BIT x)::'a) = 
+   (x = B0 \<and> iszero (number_of w::'a::{ring_char_0,number_ring}))"
+  by (simp add: iszero_def number_of_eq numeral_simps Ints_double_eq_0_iff 
+    Ints_odd_nonzero Ints_def split: bit.split)
+
+lemma iszero_number_of_0:
+  "iszero (number_of (w BIT B0) :: 'a::{ring_char_0,number_ring}) = 
+  iszero (number_of w :: 'a)"
+  by (simp only: iszero_number_of_BIT simp_thms)
+
+lemma iszero_number_of_1:
+  "~ iszero (number_of (w BIT B1)::'a::{ring_char_0,number_ring})"
+  by (simp add: iszero_number_of_BIT) 
+
+
+subsection {* The Less-Than Relation *}
+
+lemma less_number_of_eq_neg:
+  "((number_of x::'a::{ordered_idom,number_ring}) < number_of y)
+  = neg (number_of (x + uminus y) :: 'a)"
+apply (subst less_iff_diff_less_0) 
+apply (simp add: neg_def diff_minus number_of_add number_of_minus)
+done
+
+text {*
+  If @{term Numeral0} is rewritten to 0 then this rule can't be applied:
+  @{term Numeral0} IS @{term "number_of Pls"}
+*}
+
+lemma not_neg_number_of_Pls:
+  "~ neg (number_of Pls ::'a::{ordered_idom,number_ring})"
+  by (simp add: neg_def numeral_0_eq_0)
+
+lemma neg_number_of_Min:
+  "neg (number_of Min ::'a::{ordered_idom,number_ring})"
+  by (simp add: neg_def zero_less_one numeral_m1_eq_minus_1)
+
+lemma double_less_0_iff:
+  "(a + a < 0) = (a < (0::'a::ordered_idom))"
+proof -
+  have "(a + a < 0) = ((1+1)*a < 0)" by (simp add: left_distrib)
+  also have "... = (a < 0)"
+    by (simp add: mult_less_0_iff zero_less_two 
+                  order_less_not_sym [OF zero_less_two]) 
+  finally show ?thesis .
+qed
+
+lemma odd_less_0:
+  "(1 + z + z < 0) = (z < (0::int))";
+proof (cases z rule: int_cases)
+  case (nonneg n)
+  thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing
+                             le_imp_0_less [THEN order_less_imp_le])  
+next
+  case (neg n)
+  thus ?thesis by (simp del: int_Suc
+    add: int_Suc0_eq_1 [symmetric] zadd_int compare_rls)
+qed
+
+text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
+
+lemma Ints_odd_less_0: 
+  assumes in_Ints: "a \<in> Ints"
+  shows "(1 + a + a < 0) = (a < (0::'a::ordered_idom))";
+proof -
+  from in_Ints have "a \<in> range of_int" unfolding Ints_def [symmetric] .
+  then obtain z where a: "a = of_int z" ..
+  hence "((1::'a) + a + a < 0) = (of_int (1 + z + z) < (of_int 0 :: 'a))"
+    by (simp add: a)
+  also have "... = (z < 0)" by (simp only: of_int_less_iff odd_less_0)
+  also have "... = (a < 0)" by (simp add: a)
+  finally show ?thesis .
+qed
+
+lemma neg_number_of_BIT:
+  "neg (number_of (w BIT x)::'a) = 
+  neg (number_of w :: 'a::{ordered_idom,number_ring})"
+  by (simp add: neg_def number_of_eq numeral_simps double_less_0_iff
+    Ints_odd_less_0 Ints_def split: bit.split)
+
+
+text {* Less-Than or Equals *}
+
+text {* Reduces @{term "a\<le>b"} to @{term "~ (b<a)"} for ALL numerals. *}
+
+lemmas le_number_of_eq_not_less =
+  linorder_not_less [of "number_of w" "number_of v", symmetric, 
+  standard]
+
+lemma le_number_of_eq:
+    "((number_of x::'a::{ordered_idom,number_ring}) \<le> number_of y)
+     = (~ (neg (number_of (y + uminus x) :: 'a)))"
+by (simp add: le_number_of_eq_not_less less_number_of_eq_neg)
+
+
+text {* Absolute value (@{term abs}) *}
+
+lemma abs_number_of:
+  "abs(number_of x::'a::{ordered_idom,number_ring}) =
+   (if number_of x < (0::'a) then -number_of x else number_of x)"
+  by (simp add: abs_if)
+
+
+text {* Re-orientation of the equation nnn=x *}
+
+lemma number_of_reorient:
+  "(number_of w = x) = (x = number_of w)"
+  by auto
+
+
+subsection {* Simplification of arithmetic operations on integer constants. *}
+
+lemmas arith_extra_simps [standard, simp] =
+  number_of_add [symmetric]
+  number_of_minus [symmetric] numeral_m1_eq_minus_1 [symmetric]
+  number_of_mult [symmetric]
+  diff_number_of_eq abs_number_of 
+
+text {*
+  For making a minimal simpset, one must include these default simprules.
+  Also include @{text simp_thms}.
+*}
+
+lemmas arith_simps = 
+  bit.distinct
+  Pls_0_eq Min_1_eq
+  pred_Pls pred_Min pred_1 pred_0
+  succ_Pls succ_Min succ_1 succ_0
+  add_Pls add_Min add_BIT_0 add_BIT_10 add_BIT_11
+  minus_Pls minus_Min minus_1 minus_0
+  mult_Pls mult_Min mult_num1 mult_num0 
+  add_Pls_right add_Min_right
+  abs_zero abs_one arith_extra_simps
+
+text {* Simplification of relational operations *}
+
+lemmas rel_simps [simp] = 
+  eq_number_of_eq iszero_number_of_Pls nonzero_number_of_Min
+  iszero_number_of_0 iszero_number_of_1
+  less_number_of_eq_neg
+  not_neg_number_of_Pls not_neg_0 not_neg_1 not_iszero_1
+  neg_number_of_Min neg_number_of_BIT
+  le_number_of_eq
+
+
+subsection {* Simplification of arithmetic when nested to the right. *}
+
+lemma add_number_of_left [simp]:
+  "number_of v + (number_of w + z) =
+   (number_of(v + w) + z::'a::number_ring)"
+  by (simp add: add_assoc [symmetric])
+
+lemma mult_number_of_left [simp]:
+  "number_of v * (number_of w * z) =
+   (number_of(v * w) * z::'a::number_ring)"
+  by (simp add: mult_assoc [symmetric])
+
+lemma add_number_of_diff1:
+  "number_of v + (number_of w - c) = 
+  number_of(v + w) - (c::'a::number_ring)"
+  by (simp add: diff_minus add_number_of_left)
+
+lemma add_number_of_diff2 [simp]:
+  "number_of v + (c - number_of w) =
+   number_of (v + uminus w) + (c::'a::number_ring)"
+apply (subst diff_number_of_eq [symmetric])
+apply (simp only: compare_rls)
+done
+
+
+subsection {* Configuration of the code generator *}
+
+instance int :: eq ..
+
+code_datatype Pls Min Bit "number_of \<Colon> int \<Rightarrow> int"
+
+definition
+  int_aux :: "int \<Rightarrow> nat \<Rightarrow> int" where
+  "int_aux i n = (i + int n)"
+
+lemma [code]:
+  "int_aux i 0 = i"
+  "int_aux i (Suc n) = int_aux (i + 1) n" -- {* tail recursive *}
+  by (simp add: int_aux_def)+
+
+lemma [code]:
+  "int n = int_aux 0 n"
+  by (simp add: int_aux_def)
+
+definition
+  nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat" where
+  "nat_aux n i = (n + nat i)"
+
+lemma [code]: "nat_aux n i = (if i <= 0 then n else nat_aux (Suc n) (i - 1))"
+  -- {* tail recursive *}
+  by (auto simp add: nat_aux_def nat_eq_iff linorder_not_le order_less_imp_le
+    dest: zless_imp_add1_zle)
+
+lemma [code]: "nat i = nat_aux 0 i"
+  by (simp add: nat_aux_def)
+
+lemma zero_is_num_zero [code func, code inline, symmetric, normal post]:
+  "(0\<Colon>int) = number_of Numeral.Pls" 
+  by simp
+
+lemma one_is_num_one [code func, code inline, symmetric, normal post]:
+  "(1\<Colon>int) = number_of (Numeral.Pls BIT bit.B1)" 
+  by simp 
+
+code_modulename SML
+  IntDef Integer
+
+code_modulename OCaml
+  IntDef Integer
+
+code_modulename Haskell
+  IntDef Integer
+
+code_modulename SML
+  Numeral Integer
+
+code_modulename OCaml
+  Numeral Integer
+
+code_modulename Haskell
+  Numeral Integer
+
+(*FIXME: the IntInf.fromInt below hides a dependence on fixed-precision ints!*)
+
+types_code
+  "int" ("int")
+attach (term_of) {*
+val term_of_int = HOLogic.mk_number HOLogic.intT o IntInf.fromInt;
+*}
+attach (test) {*
+fun gen_int i = one_of [~1, 1] * random_range 0 i;
+*}
+
+setup {*
+let
+
+fun number_of_codegen thy defs gr dep module b (Const (@{const_name Numeral.number_of}, Type ("fun", [_, T])) $ t) =
+      if T = HOLogic.intT then
+        (SOME (fst (Codegen.invoke_tycodegen thy defs dep module false (gr, T)),
+          (Pretty.str o IntInf.toString o HOLogic.dest_numeral) t) handle TERM _ => NONE)
+      else if T = HOLogic.natT then
+        SOME (Codegen.invoke_codegen thy defs dep module b (gr,
+          Const ("IntDef.nat", HOLogic.intT --> HOLogic.natT) $
+            (Const (@{const_name Numeral.number_of}, HOLogic.intT --> HOLogic.intT) $ t)))
+      else NONE
+  | number_of_codegen _ _ _ _ _ _ _ = NONE;
+
+in
+
+Codegen.add_codegen "number_of_codegen" number_of_codegen
+
+end
+*}
+
+consts_code
+  "0 :: int"                   ("0")
+  "1 :: int"                   ("1")
+  "uminus :: int => int"       ("~")
+  "op + :: int => int => int"  ("(_ +/ _)")
+  "op * :: int => int => int"  ("(_ */ _)")
+  "op \<le> :: int => int => bool" ("(_ <=/ _)")
+  "op < :: int => int => bool" ("(_ </ _)")
+
+quickcheck_params [default_type = int]
+
+(*setup continues in theory Presburger*)
+
+hide (open) const Pls Min B0 B1 succ pred
+
+end
--- a/src/HOL/Presburger.thy	Thu May 31 18:16:51 2007 +0200
+++ b/src/HOL/Presburger.thy	Thu May 31 18:16:52 2007 +0200
@@ -9,7 +9,7 @@
 header {* Presburger Arithmetic: Cooper's Algorithm *}
 
 theory Presburger
-imports "Integ/NatSimprocs" SetInterval
+imports NatSimprocs SetInterval
 uses
   ("Tools/Presburger/cooper_dec.ML")
   ("Tools/Presburger/cooper_proof.ML")
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/int_arith1.ML	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,610 @@
+(*  Title:      HOL/int_arith1.ML
+    ID:         $Id$
+    Authors:    Larry Paulson and Tobias Nipkow
+
+Simprocs and decision procedure for linear arithmetic.
+*)
+
+(** Misc ML bindings **)
+
+val succ_Pls = thm "succ_Pls";
+val succ_Min = thm "succ_Min";
+val succ_1 = thm "succ_1";
+val succ_0 = thm "succ_0";
+
+val pred_Pls = thm "pred_Pls";
+val pred_Min = thm "pred_Min";
+val pred_1 = thm "pred_1";
+val pred_0 = thm "pred_0";
+
+val minus_Pls = thm "minus_Pls";
+val minus_Min = thm "minus_Min";
+val minus_1 = thm "minus_1";
+val minus_0 = thm "minus_0";
+
+val add_Pls = thm "add_Pls";
+val add_Min = thm "add_Min";
+val add_BIT_11 = thm "add_BIT_11";
+val add_BIT_10 = thm "add_BIT_10";
+val add_BIT_0 = thm "add_BIT_0";
+val add_Pls_right = thm "add_Pls_right";
+val add_Min_right = thm "add_Min_right";
+
+val mult_Pls = thm "mult_Pls";
+val mult_Min = thm "mult_Min";
+val mult_num1 = thm "mult_num1";
+val mult_num0 = thm "mult_num0";
+
+val neg_def = thm "neg_def";
+val iszero_def = thm "iszero_def";
+
+val number_of_succ = thm "number_of_succ";
+val number_of_pred = thm "number_of_pred";
+val number_of_minus = thm "number_of_minus";
+val number_of_add = thm "number_of_add";
+val diff_number_of_eq = thm "diff_number_of_eq";
+val number_of_mult = thm "number_of_mult";
+val double_number_of_BIT = thm "double_number_of_BIT";
+val numeral_0_eq_0 = thm "numeral_0_eq_0";
+val numeral_1_eq_1 = thm "numeral_1_eq_1";
+val numeral_m1_eq_minus_1 = thm "numeral_m1_eq_minus_1";
+val mult_minus1 = thm "mult_minus1";
+val mult_minus1_right = thm "mult_minus1_right";
+val minus_number_of_mult = thm "minus_number_of_mult";
+val zero_less_nat_eq = thm "zero_less_nat_eq";
+val eq_number_of_eq = thm "eq_number_of_eq";
+val iszero_number_of_Pls = thm "iszero_number_of_Pls";
+val nonzero_number_of_Min = thm "nonzero_number_of_Min";
+val iszero_number_of_BIT = thm "iszero_number_of_BIT";
+val iszero_number_of_0 = thm "iszero_number_of_0";
+val iszero_number_of_1 = thm "iszero_number_of_1";
+val less_number_of_eq_neg = thm "less_number_of_eq_neg";
+val le_number_of_eq = thm "le_number_of_eq";
+val not_neg_number_of_Pls = thm "not_neg_number_of_Pls";
+val neg_number_of_Min = thm "neg_number_of_Min";
+val neg_number_of_BIT = thm "neg_number_of_BIT";
+val le_number_of_eq_not_less = thm "le_number_of_eq_not_less";
+val abs_number_of = thm "abs_number_of";
+val number_of_reorient = thm "number_of_reorient";
+val add_number_of_left = thm "add_number_of_left";
+val mult_number_of_left = thm "mult_number_of_left";
+val add_number_of_diff1 = thm "add_number_of_diff1";
+val add_number_of_diff2 = thm "add_number_of_diff2";
+val less_iff_diff_less_0 = thm "less_iff_diff_less_0";
+val eq_iff_diff_eq_0 = thm "eq_iff_diff_eq_0";
+val le_iff_diff_le_0 = thm "le_iff_diff_le_0";
+
+val arith_extra_simps = thms "arith_extra_simps";
+val arith_simps = thms "arith_simps";
+val rel_simps = thms "rel_simps";
+
+val zless_imp_add1_zle = thm "zless_imp_add1_zle";
+
+val combine_common_factor = thm "combine_common_factor";
+val eq_add_iff1 = thm "eq_add_iff1";
+val eq_add_iff2 = thm "eq_add_iff2";
+val less_add_iff1 = thm "less_add_iff1";
+val less_add_iff2 = thm "less_add_iff2";
+val le_add_iff1 = thm "le_add_iff1";
+val le_add_iff2 = thm "le_add_iff2";
+
+val arith_special = thms "arith_special";
+
+structure Int_Numeral_Base_Simprocs =
+  struct
+  fun prove_conv tacs ctxt (_: thm list) (t, u) =
+    if t aconv u then NONE
+    else
+      let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))
+      in SOME (Goal.prove ctxt [] [] eq (K (EVERY tacs))) end
+
+  fun prove_conv_nohyps tacs sg = prove_conv tacs sg [];
+
+  fun prep_simproc (name, pats, proc) =
+    Simplifier.simproc (the_context()) name pats proc;
+
+  fun is_numeral (Const(@{const_name Numeral.number_of}, _) $ w) = true
+    | is_numeral _ = false
+
+  fun simplify_meta_eq f_number_of_eq f_eq =
+      mk_meta_eq ([f_eq, f_number_of_eq] MRS trans)
+
+  (*reorientation simprules using ==, for the following simproc*)
+  val meta_zero_reorient = zero_reorient RS eq_reflection
+  val meta_one_reorient = one_reorient RS eq_reflection
+  val meta_number_of_reorient = number_of_reorient RS eq_reflection
+
+  (*reorientation simplification procedure: reorients (polymorphic) 
+    0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
+  fun reorient_proc sg _ (_ $ t $ u) =
+    case u of
+	Const(@{const_name HOL.zero}, _) => NONE
+      | Const(@{const_name HOL.one}, _) => NONE
+      | Const(@{const_name Numeral.number_of}, _) $ _ => NONE
+      | _ => SOME (case t of
+		  Const(@{const_name HOL.zero}, _) => meta_zero_reorient
+		| Const(@{const_name HOL.one}, _) => meta_one_reorient
+		| Const(@{const_name Numeral.number_of}, _) $ _ => meta_number_of_reorient)
+
+  val reorient_simproc = 
+      prep_simproc ("reorient_simproc", ["0=x", "1=x", "number_of w = x"], reorient_proc)
+
+  end;
+
+
+Addsimprocs [Int_Numeral_Base_Simprocs.reorient_simproc];
+
+
+structure Int_Numeral_Simprocs =
+struct
+
+(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in Int_Numeral_Base_Simprocs
+  isn't complicated by the abstract 0 and 1.*)
+val numeral_syms = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym];
+
+(** New term ordering so that AC-rewriting brings numerals to the front **)
+
+(*Order integers by absolute value and then by sign. The standard integer
+  ordering is not well-founded.*)
+fun num_ord (i,j) =
+      (case IntInf.compare (IntInf.abs i, IntInf.abs j) of
+            EQUAL => int_ord (IntInf.sign i, IntInf.sign j) 
+          | ord => ord);
+
+(*This resembles Term.term_ord, but it puts binary numerals before other
+  non-atomic terms.*)
+local open Term 
+in 
+fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) =
+      (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
+  | numterm_ord
+     (Const(@{const_name Numeral.number_of}, _) $ v, Const(@{const_name Numeral.number_of}, _) $ w) =
+     num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
+  | numterm_ord (Const(@{const_name Numeral.number_of}, _) $ _, _) = LESS
+  | numterm_ord (_, Const(@{const_name Numeral.number_of}, _) $ _) = GREATER
+  | numterm_ord (t, u) =
+      (case int_ord (size_of_term t, size_of_term u) of
+        EQUAL =>
+          let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
+            (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord)
+          end
+      | ord => ord)
+and numterms_ord (ts, us) = list_ord numterm_ord (ts, us)
+end;
+
+fun numtermless tu = (numterm_ord tu = LESS);
+
+(*Defined in this file, but perhaps needed only for Int_Numeral_Base_Simprocs of type nat.*)
+val num_ss = HOL_ss settermless numtermless;
+
+
+(** Utilities **)
+
+fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n;
+
+fun find_first_numeral past (t::terms) =
+        ((snd (HOLogic.dest_number t), rev past @ terms)
+         handle TERM _ => find_first_numeral (t::past) terms)
+  | find_first_numeral past [] = raise TERM("find_first_numeral", []);
+
+val mk_plus = HOLogic.mk_binop @{const_name HOL.plus};
+
+fun mk_minus t = 
+  let val T = Term.fastype_of t
+  in Const (@{const_name HOL.uminus}, T --> T) $ t
+  end;
+
+(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
+fun mk_sum T []        = mk_number T 0
+  | mk_sum T [t,u]     = mk_plus (t, u)
+  | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
+
+(*this version ALWAYS includes a trailing zero*)
+fun long_mk_sum T []        = mk_number T 0
+  | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
+
+val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT;
+
+(*decompose additions AND subtractions as a sum*)
+fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (pos, u, ts))
+  | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) =
+        dest_summing (pos, t, dest_summing (not pos, u, ts))
+  | dest_summing (pos, t, ts) =
+        if pos then t::ts else mk_minus t :: ts;
+
+fun dest_sum t = dest_summing (true, t, []);
+
+val mk_diff = HOLogic.mk_binop @{const_name HOL.minus};
+val dest_diff = HOLogic.dest_bin @{const_name HOL.minus} Term.dummyT;
+
+val mk_times = HOLogic.mk_binop @{const_name HOL.times};
+
+fun mk_prod T = 
+  let val one = mk_number T 1
+  fun mk [] = one
+    | mk [t] = t
+    | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts)
+  in mk end;
+
+(*This version ALWAYS includes a trailing one*)
+fun long_mk_prod T []        = mk_number T 1
+  | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts);
+
+val dest_times = HOLogic.dest_bin @{const_name HOL.times} Term.dummyT;
+
+fun dest_prod t =
+      let val (t,u) = dest_times t
+      in  dest_prod t @ dest_prod u  end
+      handle TERM _ => [t];
+
+(*DON'T do the obvious simplifications; that would create special cases*)
+fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t);
+
+(*Express t as a product of (possibly) a numeral with other sorted terms*)
+fun dest_coeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_coeff (~sign) t
+  | dest_coeff sign t =
+    let val ts = sort Term.term_ord (dest_prod t)
+        val (n, ts') = find_first_numeral [] ts
+                          handle TERM _ => (1, ts)
+    in (sign*n, mk_prod (Term.fastype_of t) ts') end;
+
+(*Find first coefficient-term THAT MATCHES u*)
+fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
+  | find_first_coeff past u (t::terms) =
+        let val (n,u') = dest_coeff 1 t
+        in  if u aconv u' then (n, rev past @ terms)
+                          else find_first_coeff (t::past) u terms
+        end
+        handle TERM _ => find_first_coeff (t::past) u terms;
+
+(*Fractions as pairs of ints. Can't use Rat.rat because the representation
+  needs to preserve negative values in the denominator.*)
+fun mk_frac (p, q : IntInf.int) = if q = 0 then raise Div else (p, q);
+
+(*Don't reduce fractions; sums must be proved by rule add_frac_eq.
+  Fractions are reduced later by the cancel_numeral_factor simproc.*)
+fun add_frac ((p1 : IntInf.int, q1 : IntInf.int), (p2, q2)) = (p1 * q2 + p2 * q1, q1 * q2);
+
+val mk_divide = HOLogic.mk_binop @{const_name HOL.divide};
+
+(*Build term (p / q) * t*)
+fun mk_fcoeff ((p, q), t) =
+  let val T = Term.fastype_of t
+  in  mk_times (mk_divide (mk_number T p, mk_number T q), t) end;
+
+(*Express t as a product of a fraction with other sorted terms*)
+fun dest_fcoeff sign (Const (@{const_name HOL.uminus}, _) $ t) = dest_fcoeff (~sign) t
+  | dest_fcoeff sign (Const (@{const_name HOL.divide}, _) $ t $ u) =
+    let val (p, t') = dest_coeff sign t
+        val (q, u') = dest_coeff 1 u
+    in  (mk_frac (p, q), mk_divide (t', u')) end
+  | dest_fcoeff sign t =
+    let val (p, t') = dest_coeff sign t
+        val T = Term.fastype_of t
+    in  (mk_frac (p, 1), mk_divide (t', mk_number T 1)) end;
+
+
+(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
+val add_0s =  thms "add_0s";
+val mult_1s = thms "mult_1s";
+
+(*Simplify inverse Numeral1, a/Numeral1*)
+val inverse_1s = [@{thm inverse_numeral_1}];
+val divide_1s = [@{thm divide_numeral_1}];
+
+(*To perform binary arithmetic.  The "left" rewriting handles patterns
+  created by the Int_Numeral_Base_Simprocs, such as 3 * (5 * x). *)
+val simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym,
+                 add_number_of_left, mult_number_of_left] @
+                arith_simps @ rel_simps;
+
+(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
+  during re-arrangement*)
+val non_add_simps =
+  subtract Thm.eq_thm [add_number_of_left, number_of_add RS sym] simps;
+
+(*To evaluate binary negations of coefficients*)
+val minus_simps = [numeral_m1_eq_minus_1 RS sym, number_of_minus RS sym,
+                   minus_1, minus_0, minus_Pls, minus_Min,
+                   pred_1, pred_0, pred_Pls, pred_Min];
+
+(*To let us treat subtraction as addition*)
+val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
+
+(*To let us treat division as multiplication*)
+val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
+
+(*push the unary minus down: - x * y = x * - y *)
+val minus_mult_eq_1_to_2 =
+    [@{thm minus_mult_left} RS sym, @{thm minus_mult_right}] MRS trans |> standard;
+
+(*to extract again any uncancelled minuses*)
+val minus_from_mult_simps =
+    [@{thm minus_minus}, @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym];
+
+(*combine unary minus with numeric literals, however nested within a product*)
+val mult_minus_simps =
+    [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2];
+
+(*Apply the given rewrite (if present) just once*)
+fun trans_tac NONE      = all_tac
+  | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans));
+
+fun simplify_meta_eq rules =
+  let val ss0 = HOL_basic_ss addeqcongs [eq_cong2] addsimps rules
+  in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end
+
+structure CancelNumeralsCommon =
+  struct
+  val mk_sum            = mk_sum
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val find_first_coeff  = find_first_coeff []
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
+    diff_simps @ minus_simps @ add_ac
+  val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
+  val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
+
+  val numeral_simp_ss = HOL_ss addsimps add_0s @ simps
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s)
+  end;
+
+
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val bal_add1 = eq_add_iff1 RS trans
+  val bal_add2 = eq_add_iff2 RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
+  val bal_add1 = less_add_iff1 RS trans
+  val bal_add2 = less_add_iff2 RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
+  val bal_add1 = le_add_iff1 RS trans
+  val bal_add2 = le_add_iff2 RS trans
+);
+
+val cancel_numerals =
+  map Int_Numeral_Base_Simprocs.prep_simproc
+   [("inteq_cancel_numerals",
+     ["(l::'a::number_ring) + m = n",
+      "(l::'a::number_ring) = m + n",
+      "(l::'a::number_ring) - m = n",
+      "(l::'a::number_ring) = m - n",
+      "(l::'a::number_ring) * m = n",
+      "(l::'a::number_ring) = m * n"],
+     K EqCancelNumerals.proc),
+    ("intless_cancel_numerals",
+     ["(l::'a::{ordered_idom,number_ring}) + m < n",
+      "(l::'a::{ordered_idom,number_ring}) < m + n",
+      "(l::'a::{ordered_idom,number_ring}) - m < n",
+      "(l::'a::{ordered_idom,number_ring}) < m - n",
+      "(l::'a::{ordered_idom,number_ring}) * m < n",
+      "(l::'a::{ordered_idom,number_ring}) < m * n"],
+     K LessCancelNumerals.proc),
+    ("intle_cancel_numerals",
+     ["(l::'a::{ordered_idom,number_ring}) + m <= n",
+      "(l::'a::{ordered_idom,number_ring}) <= m + n",
+      "(l::'a::{ordered_idom,number_ring}) - m <= n",
+      "(l::'a::{ordered_idom,number_ring}) <= m - n",
+      "(l::'a::{ordered_idom,number_ring}) * m <= n",
+      "(l::'a::{ordered_idom,number_ring}) <= m * n"],
+     K LeCancelNumerals.proc)];
+
+
+structure CombineNumeralsData =
+  struct
+  type coeff            = IntInf.int
+  val iszero            = (fn x : IntInf.int => x = 0)
+  val add               = IntInf.+
+  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val left_distrib      = combine_common_factor RS trans
+  val prove_conv        = Int_Numeral_Base_Simprocs.prove_conv_nohyps
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
+    diff_simps @ minus_simps @ add_ac
+  val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
+  val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
+
+  val numeral_simp_ss = HOL_ss addsimps add_0s @ simps
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s)
+  end;
+
+structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
+
+(*Version for fields, where coefficients can be fractions*)
+structure FieldCombineNumeralsData =
+  struct
+  type coeff            = IntInf.int * IntInf.int
+  val iszero            = (fn (p : IntInf.int, q) => p = 0)
+  val add               = add_frac
+  val mk_sum            = long_mk_sum
+  val dest_sum          = dest_sum
+  val mk_coeff          = mk_fcoeff
+  val dest_coeff        = dest_fcoeff 1
+  val left_distrib      = combine_common_factor RS trans
+  val prove_conv        = Int_Numeral_Base_Simprocs.prove_conv_nohyps
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
+    inverse_1s @ divide_simps @ diff_simps @ minus_simps @ add_ac
+  val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
+  val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
+
+  val numeral_simp_ss = HOL_ss addsimps add_0s @ simps @ [@{thm add_frac_eq}]
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s @ divide_1s)
+  end;
+
+structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData);
+
+val combine_numerals =
+  Int_Numeral_Base_Simprocs.prep_simproc
+    ("int_combine_numerals", 
+     ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], 
+     K CombineNumerals.proc);
+
+val field_combine_numerals =
+  Int_Numeral_Base_Simprocs.prep_simproc
+    ("field_combine_numerals", 
+     ["(i::'a::{number_ring,field,division_by_zero}) + j",
+      "(i::'a::{number_ring,field,division_by_zero}) - j"], 
+     K FieldCombineNumerals.proc);
+
+end;
+
+Addsimprocs Int_Numeral_Simprocs.cancel_numerals;
+Addsimprocs [Int_Numeral_Simprocs.combine_numerals];
+Addsimprocs [Int_Numeral_Simprocs.field_combine_numerals];
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s, by (Simp_tac 1));
+
+test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)";
+
+test "2*u = (u::int)";
+test "(i + j + 12 + (k::int)) - 15 = y";
+test "(i + j + 12 + (k::int)) - 5 = y";
+
+test "y - b < (b::int)";
+test "y - (3*b + c) < (b::int) - 2*c";
+
+test "(2*x - (u*v) + y) - v*3*u = (w::int)";
+test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)";
+test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)";
+test "u*v - (x*u*v + (u*v)*4 + y) = (w::int)";
+
+test "(i + j + 12 + (k::int)) = u + 15 + y";
+test "(i + j*2 + 12 + (k::int)) = j + 5 + y";
+
+test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)";
+
+test "a + -(b+c) + b = (d::int)";
+test "a + -(b+c) - b = (d::int)";
+
+(*negative numerals*)
+test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz";
+test "(i + j + -3 + (k::int)) < u + 5 + y";
+test "(i + j + 3 + (k::int)) < u + -6 + y";
+test "(i + j + -12 + (k::int)) - 15 = y";
+test "(i + j + 12 + (k::int)) - -15 = y";
+test "(i + j + -12 + (k::int)) - -15 = y";
+*)
+
+
+(** Constant folding for multiplication in semirings **)
+
+(*We do not need folding for addition: combine_numerals does the same thing*)
+
+structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA =
+struct
+  val assoc_ss = HOL_ss addsimps mult_ac
+  val eq_reflection = eq_reflection
+end;
+
+structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
+
+val assoc_fold_simproc =
+  Int_Numeral_Base_Simprocs.prep_simproc
+   ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
+    K Semiring_Times_Assoc.proc);
+
+Addsimprocs [assoc_fold_simproc];
+
+
+
+
+(*** decision procedure for linear arithmetic ***)
+
+(*---------------------------------------------------------------------------*)
+(* Linear arithmetic                                                         *)
+(*---------------------------------------------------------------------------*)
+
+(*
+Instantiation of the generic linear arithmetic package for int.
+*)
+
+(* Update parameters of arithmetic prover *)
+local
+
+(* reduce contradictory <= to False *)
+val add_rules =
+    simp_thms @ arith_simps @ rel_simps @ arith_special @
+    [@{thm neg_le_iff_le}, @{thm numeral_0_eq_0}, @{thm numeral_1_eq_1},
+     @{thm minus_zero}, @{thm diff_minus}, @{thm left_minus}, @{thm right_minus},
+     @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_num1}, @{thm mult_1_right},
+     @{thm minus_mult_left} RS sym, @{thm minus_mult_right} RS sym,
+     @{thm minus_add_distrib}, @{thm minus_minus}, @{thm mult_assoc},
+     of_nat_0, of_nat_1, of_nat_Suc, of_nat_add, of_nat_mult,
+     of_int_0, of_int_1, of_int_add, of_int_mult, int_eq_of_nat]
+
+val nat_inj_thms = [zle_int RS iffD2, int_int_eq RS iffD2]
+
+val Int_Numeral_Base_Simprocs = assoc_fold_simproc
+  :: Int_Numeral_Simprocs.combine_numerals
+  :: Int_Numeral_Simprocs.cancel_numerals;
+
+in
+
+val int_arith_setup =
+  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+   {add_mono_thms = add_mono_thms,
+    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms,
+    inj_thms = nat_inj_thms @ inj_thms,
+    lessD = lessD @ [zless_imp_add1_zle],
+    neqE = neqE,
+    simpset = simpset addsimps add_rules
+                      addsimprocs Int_Numeral_Base_Simprocs
+                      addcongs [if_weak_cong]}) #>
+  arith_inj_const ("Nat.of_nat", HOLogic.natT --> HOLogic.intT) #>
+  arith_inj_const ("IntDef.int", HOLogic.natT --> HOLogic.intT) #>
+  arith_discrete "IntDef.int"
+
+end;
+
+val fast_int_arith_simproc =
+  Simplifier.simproc @{theory}
+  "fast_int_arith" 
+     ["(m::'a::{ordered_idom,number_ring}) < n",
+      "(m::'a::{ordered_idom,number_ring}) <= n",
+      "(m::'a::{ordered_idom,number_ring}) = n"] Fast_Arith.lin_arith_prover;
+
+Addsimprocs [fast_int_arith_simproc];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/int_factor_simprocs.ML	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,335 @@
+(*  Title:      HOL/int_factor_simprocs.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   2000  University of Cambridge
+
+Factor cancellation simprocs for the integers (and for fields).
+
+This file can't be combined with int_arith1 because it requires IntDiv.thy.
+*)
+
+
+(*To quote from Provers/Arith/cancel_numeral_factor.ML:
+
+Cancels common coefficients in balanced expressions:
+
+     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
+
+where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
+and d = gcd(m,m') and n=m/d and n'=m'/d.
+*)
+
+val rel_number_of = [eq_number_of_eq, less_number_of_eq_neg, le_number_of_eq];
+
+(** Factor cancellation theorems for integer division (div, not /) **)
+
+Goal "!!k::int. k~=0 ==> (k*m) div (k*n) = (m div n)";
+by (stac @{thm zdiv_zmult_zmult1} 1);
+by Auto_tac;
+qed "int_mult_div_cancel1";
+
+(*For ExtractCommonTermFun, cancelling common factors*)
+Goal "(k*m) div (k*n) = (if k = (0::int) then 0 else m div n)";
+by (simp_tac (simpset() addsimps [int_mult_div_cancel1]) 1);
+qed "int_mult_div_cancel_disj";
+
+
+local
+  open Int_Numeral_Simprocs
+in
+
+structure CancelNumeralFactorCommon =
+  struct
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = HOL_ss addsimps minus_from_mult_simps @ mult_1s
+  val norm_ss2 = HOL_ss addsimps simps @ mult_minus_simps
+  val norm_ss3 = HOL_ss addsimps mult_ac
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3))
+
+  val numeral_simp_ss = HOL_ss addsimps rel_number_of @ simps
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq = Int_Numeral_Simprocs.simplify_meta_eq
+    [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left},
+      @{thm mult_zero_right}, @{thm mult_num1}, @{thm mult_1_right}];
+  end
+
+(*Version for integer division*)
+structure IntDivCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT
+  val cancel = int_mult_div_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+(*Version for fields*)
+structure DivideCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name HOL.divide}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT
+  val cancel = @{thm mult_divide_cancel_left} RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val cancel = @{thm mult_cancel_left} RS trans
+  val neg_exchanges = false
+)
+
+structure LessCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
+  val cancel = @{thm mult_less_cancel_left} RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
+  val cancel = @{thm mult_le_cancel_left} RS trans
+  val neg_exchanges = true
+)
+
+val cancel_numeral_factors =
+  map Int_Numeral_Base_Simprocs.prep_simproc
+   [("ring_eq_cancel_numeral_factor",
+     ["(l::'a::{idom,number_ring}) * m = n",
+      "(l::'a::{idom,number_ring}) = m * n"],
+     K EqCancelNumeralFactor.proc),
+    ("ring_less_cancel_numeral_factor",
+     ["(l::'a::{ordered_idom,number_ring}) * m < n",
+      "(l::'a::{ordered_idom,number_ring}) < m * n"],
+     K LessCancelNumeralFactor.proc),
+    ("ring_le_cancel_numeral_factor",
+     ["(l::'a::{ordered_idom,number_ring}) * m <= n",
+      "(l::'a::{ordered_idom,number_ring}) <= m * n"],
+     K LeCancelNumeralFactor.proc),
+    ("int_div_cancel_numeral_factors",
+     ["((l::int) * m) div n", "(l::int) div (m * n)"],
+     K IntDivCancelNumeralFactor.proc),
+    ("divide_cancel_numeral_factor",
+     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
+      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
+      "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
+     K DivideCancelNumeralFactor.proc)];
+
+(* referenced by rat_arith.ML *)
+val field_cancel_numeral_factors =
+  map Int_Numeral_Base_Simprocs.prep_simproc
+   [("field_eq_cancel_numeral_factor",
+     ["(l::'a::{field,number_ring}) * m = n",
+      "(l::'a::{field,number_ring}) = m * n"],
+     K EqCancelNumeralFactor.proc),
+    ("field_cancel_numeral_factor",
+     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
+      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
+      "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
+     K DivideCancelNumeralFactor.proc)]
+
+end;
+
+Addsimprocs cancel_numeral_factors;
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1));
+
+test "9*x = 12 * (y::int)";
+test "(9*x) div (12 * (y::int)) = z";
+test "9*x < 12 * (y::int)";
+test "9*x <= 12 * (y::int)";
+
+test "-99*x = 132 * (y::int)";
+test "(-99*x) div (132 * (y::int)) = z";
+test "-99*x < 132 * (y::int)";
+test "-99*x <= 132 * (y::int)";
+
+test "999*x = -396 * (y::int)";
+test "(999*x) div (-396 * (y::int)) = z";
+test "999*x < -396 * (y::int)";
+test "999*x <= -396 * (y::int)";
+
+test "-99*x = -81 * (y::int)";
+test "(-99*x) div (-81 * (y::int)) = z";
+test "-99*x <= -81 * (y::int)";
+test "-99*x < -81 * (y::int)";
+
+test "-2 * x = -1 * (y::int)";
+test "-2 * x = -(y::int)";
+test "(-2 * x) div (-1 * (y::int)) = z";
+test "-2 * x < -(y::int)";
+test "-2 * x <= -1 * (y::int)";
+test "-x < -23 * (y::int)";
+test "-x <= -23 * (y::int)";
+*)
+
+(*And the same examples for fields such as rat or real:
+test "0 <= (y::rat) * -2";
+test "9*x = 12 * (y::rat)";
+test "(9*x) / (12 * (y::rat)) = z";
+test "9*x < 12 * (y::rat)";
+test "9*x <= 12 * (y::rat)";
+
+test "-99*x = 132 * (y::rat)";
+test "(-99*x) / (132 * (y::rat)) = z";
+test "-99*x < 132 * (y::rat)";
+test "-99*x <= 132 * (y::rat)";
+
+test "999*x = -396 * (y::rat)";
+test "(999*x) / (-396 * (y::rat)) = z";
+test "999*x < -396 * (y::rat)";
+test "999*x <= -396 * (y::rat)";
+
+test  "(- ((2::rat) * x) <= 2 * y)";
+test "-99*x = -81 * (y::rat)";
+test "(-99*x) / (-81 * (y::rat)) = z";
+test "-99*x <= -81 * (y::rat)";
+test "-99*x < -81 * (y::rat)";
+
+test "-2 * x = -1 * (y::rat)";
+test "-2 * x = -(y::rat)";
+test "(-2 * x) / (-1 * (y::rat)) = z";
+test "-2 * x < -(y::rat)";
+test "-2 * x <= -1 * (y::rat)";
+test "-x < -23 * (y::rat)";
+test "-x <= -23 * (y::rat)";
+*)
+
+
+(** Declarations for ExtractCommonTerm **)
+
+local
+  open Int_Numeral_Simprocs
+in
+
+(*Find first term that matches u*)
+fun find_first_t past u []         = raise TERM ("find_first_t", [])
+  | find_first_t past u (t::terms) =
+        if u aconv t then (rev past @ terms)
+        else find_first_t (t::past) u terms
+        handle TERM _ => find_first_t (t::past) u terms;
+
+(** Final simplification for the CancelFactor simprocs **)
+val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq  
+  [@{thm mult_1_left}, mult_1_right, @{thm zdiv_1}, numeral_1_eq_1];
+
+fun cancel_simplify_meta_eq cancel_th ss th =
+    simplify_one ss (([th, cancel_th]) MRS trans);
+
+(*At present, long_mk_prod creates Numeral1, so this requires the axclass
+  number_ring*)
+structure CancelFactorCommon =
+  struct
+  val mk_sum            = long_mk_prod
+  val dest_sum          = dest_prod
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff
+  val find_first        = find_first_t []
+  val trans_tac         = fn _ => trans_tac
+  val norm_ss = HOL_ss addsimps mult_1s @ mult_ac
+  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
+  end;
+
+(*mult_cancel_left requires a ring with no zero divisors.*)
+structure EqCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_cancel_left}
+);
+
+(*int_mult_div_cancel_disj is for integer division (div).*)
+structure IntDivCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT
+  val simplify_meta_eq  = cancel_simplify_meta_eq int_mult_div_cancel_disj
+);
+
+(*Version for all fields, including unordered ones (type complex).*)
+structure DivideCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name HOL.divide}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.divide} Term.dummyT
+  val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_divide_cancel_eq_if}
+);
+
+(*The number_ring class is necessary because the simprocs refer to the
+  binary number 1.  FIXME: probably they could use 1 instead.*)
+val cancel_factors =
+  map Int_Numeral_Base_Simprocs.prep_simproc
+   [("ring_eq_cancel_factor",
+     ["(l::'a::{idom,number_ring}) * m = n",
+      "(l::'a::{idom,number_ring}) = m * n"],
+    K EqCancelFactor.proc),
+    ("int_div_cancel_factor",
+     ["((l::int) * m) div n", "(l::int) div (m * n)"],
+     K IntDivCancelFactor.proc),
+    ("divide_cancel_factor",
+     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
+      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)"],
+     K DivideCancelFactor.proc)];
+
+end;
+
+Addsimprocs cancel_factors;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Asm_simp_tac 1));
+
+test "x*k = k*(y::int)";
+test "k = k*(y::int)";
+test "a*(b*c) = (b::int)";
+test "a*(b*c) = d*(b::int)*(x*a)";
+
+test "(x*k) div (k*(y::int)) = (uu::int)";
+test "(k) div (k*(y::int)) = (uu::int)";
+test "(a*(b*c)) div ((b::int)) = (uu::int)";
+test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)";
+*)
+
+(*And the same examples for fields such as rat or real:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Asm_simp_tac 1));
+
+test "x*k = k*(y::rat)";
+test "k = k*(y::rat)";
+test "a*(b*c) = (b::rat)";
+test "a*(b*c) = d*(b::rat)*(x*a)";
+
+
+test "(x*k) / (k*(y::rat)) = (uu::rat)";
+test "(k) / (k*(y::rat)) = (uu::rat)";
+test "(a*(b*c)) / ((b::rat)) = (uu::rat)";
+test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)";
+
+(*FIXME: what do we do about this?*)
+test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z";
+*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/nat_simprocs.ML	Thu May 31 18:16:52 2007 +0200
@@ -0,0 +1,550 @@
+(*  Title:      HOL/nat_simprocs.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   2000  University of Cambridge
+
+Simprocs for nat numerals.
+*)
+
+structure Nat_Numeral_Simprocs =
+struct
+
+(*Maps n to #n for n = 0, 1, 2*)
+val numeral_syms =
+       [nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym, numeral_2_eq_2 RS sym];
+val numeral_sym_ss = HOL_ss addsimps numeral_syms;
+
+fun rename_numerals th =
+    simplify numeral_sym_ss (Thm.transfer (the_context ()) th);
+
+(*Utilities*)
+
+fun mk_number n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_numeral n;
+fun dest_number t = IntInf.max (0, snd (HOLogic.dest_number t));
+
+fun find_first_numeral past (t::terms) =
+        ((dest_number t, t, rev past @ terms)
+         handle TERM _ => find_first_numeral (t::past) terms)
+  | find_first_numeral past [] = raise TERM("find_first_numeral", []);
+
+val zero = mk_number 0;
+val mk_plus = HOLogic.mk_binop @{const_name HOL.plus};
+
+(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
+fun mk_sum []        = zero
+  | mk_sum [t,u]     = mk_plus (t, u)
+  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+(*this version ALWAYS includes a trailing zero*)
+fun long_mk_sum []        = HOLogic.zero
+  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT;
+
+
+(** Other simproc items **)
+
+val trans_tac = Int_Numeral_Simprocs.trans_tac;
+
+val bin_simps =
+     [nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym,
+      add_nat_number_of, nat_number_of_add_left, 
+      diff_nat_number_of, le_number_of_eq_not_less,
+      mult_nat_number_of, nat_number_of_mult_left, 
+      less_nat_number_of, 
+      @{thm Let_number_of}, nat_number_of] @
+     arith_simps @ rel_simps;
+
+fun prep_simproc (name, pats, proc) =
+  Simplifier.simproc (the_context ()) name pats proc;
+
+
+(*** CancelNumerals simprocs ***)
+
+val one = mk_number 1;
+val mk_times = HOLogic.mk_binop @{const_name HOL.times};
+
+fun mk_prod [] = one
+  | mk_prod [t] = t
+  | mk_prod (t :: ts) = if t = one then mk_prod ts
+                        else mk_times (t, mk_prod ts);
+
+val dest_times = HOLogic.dest_bin @{const_name HOL.times} HOLogic.natT;
+
+fun dest_prod t =
+      let val (t,u) = dest_times t
+      in  dest_prod t @ dest_prod u  end
+      handle TERM _ => [t];
+
+(*DON'T do the obvious simplifications; that would create special cases*)
+fun mk_coeff (k,t) = mk_times (mk_number k, t);
+
+(*Express t as a product of (possibly) a numeral with other factors, sorted*)
+fun dest_coeff t =
+    let val ts = sort Term.term_ord (dest_prod t)
+        val (n, _, ts') = find_first_numeral [] ts
+                          handle TERM _ => (1, one, ts)
+    in (n, mk_prod ts') end;
+
+(*Find first coefficient-term THAT MATCHES u*)
+fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
+  | find_first_coeff past u (t::terms) =
+        let val (n,u') = dest_coeff t
+        in  if u aconv u' then (n, rev past @ terms)
+                          else find_first_coeff (t::past) u terms
+        end
+        handle TERM _ => find_first_coeff (t::past) u terms;
+
+
+(*Split up a sum into the list of its constituent terms, on the way removing any
+  Sucs and counting them.*)
+fun dest_Suc_sum (Const ("Suc", _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts))
+  | dest_Suc_sum (t, (k,ts)) = 
+      let val (t1,t2) = dest_plus t
+      in  dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts)))  end
+      handle TERM _ => (k, t::ts);
+
+(*Code for testing whether numerals are already used in the goal*)
+fun is_numeral (Const(@{const_name Numeral.number_of}, _) $ w) = true
+  | is_numeral _ = false;
+
+fun prod_has_numeral t = exists is_numeral (dest_prod t);
+
+(*The Sucs found in the term are converted to a binary numeral. If relaxed is false,
+  an exception is raised unless the original expression contains at least one
+  numeral in a coefficient position.  This prevents nat_combine_numerals from 
+  introducing numerals to goals.*)
+fun dest_Sucs_sum relaxed t = 
+  let val (k,ts) = dest_Suc_sum (t,(0,[]))
+  in
+     if relaxed orelse exists prod_has_numeral ts then 
+       if k=0 then ts
+       else mk_number (IntInf.fromInt k) :: ts
+     else raise TERM("Nat_Numeral_Simprocs.dest_Sucs_sum", [t])
+  end;
+
+
+(*Simplify 1*n and n*1 to n*)
+val add_0s  = map rename_numerals [add_0, add_0_right];
+val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}];
+
+(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*)
+
+(*And these help the simproc return False when appropriate, which helps
+  the arith prover.*)
+val contra_rules = [add_Suc, add_Suc_right, Zero_not_Suc, Suc_not_Zero,
+                    le_0_eq];
+
+val simplify_meta_eq =
+    Int_Numeral_Simprocs.simplify_meta_eq
+        ([nat_numeral_0_eq_0, numeral_1_eq_Suc_0, add_0, add_0_right,
+          mult_0, mult_0_right, mult_1, mult_1_right] @ contra_rules);
+
+
+(*Like HOL_ss but with an ordering that brings numerals to the front
+  under AC-rewriting.*)
+val num_ss = Int_Numeral_Simprocs.num_ss;
+
+(*** Applying CancelNumeralsFun ***)
+
+structure CancelNumeralsCommon =
+  struct
+  val mk_sum            = (fn T:typ => mk_sum)
+  val dest_sum          = dest_Sucs_sum true
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff
+  val find_first_coeff  = find_first_coeff []
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
+    [Suc_eq_add_numeral_1_left] @ add_ac
+  val norm_ss2 = num_ss addsimps bin_simps @ add_ac @ mult_ac
+  fun norm_tac ss = 
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+
+  val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps;
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss));
+  val simplify_meta_eq  = simplify_meta_eq
+  end;
+
+
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val bal_add1 = nat_eq_add_iff1 RS trans
+  val bal_add2 = nat_eq_add_iff2 RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
+  val bal_add1 = nat_less_add_iff1 RS trans
+  val bal_add2 = nat_less_add_iff2 RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
+  val bal_add1 = nat_le_add_iff1 RS trans
+  val bal_add2 = nat_le_add_iff2 RS trans
+);
+
+structure DiffCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name HOL.minus}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT
+  val bal_add1 = nat_diff_add_eq1 RS trans
+  val bal_add2 = nat_diff_add_eq2 RS trans
+);
+
+
+val cancel_numerals =
+  map prep_simproc
+   [("nateq_cancel_numerals",
+     ["(l::nat) + m = n", "(l::nat) = m + n",
+      "(l::nat) * m = n", "(l::nat) = m * n",
+      "Suc m = n", "m = Suc n"],
+     K EqCancelNumerals.proc),
+    ("natless_cancel_numerals",
+     ["(l::nat) + m < n", "(l::nat) < m + n",
+      "(l::nat) * m < n", "(l::nat) < m * n",
+      "Suc m < n", "m < Suc n"],
+     K LessCancelNumerals.proc),
+    ("natle_cancel_numerals",
+     ["(l::nat) + m <= n", "(l::nat) <= m + n",
+      "(l::nat) * m <= n", "(l::nat) <= m * n",
+      "Suc m <= n", "m <= Suc n"],
+     K LeCancelNumerals.proc),
+    ("natdiff_cancel_numerals",
+     ["((l::nat) + m) - n", "(l::nat) - (m + n)",
+      "(l::nat) * m - n", "(l::nat) - m * n",
+      "Suc m - n", "m - Suc n"],
+     K DiffCancelNumerals.proc)];
+
+
+(*** Applying CombineNumeralsFun ***)
+
+structure CombineNumeralsData =
+  struct
+  type coeff            = IntInf.int
+  val iszero            = (fn x : IntInf.int => x = 0)
+  val add               = IntInf.+
+  val mk_sum            = (fn T:typ => long_mk_sum)  (*to work for 2*x + 3*x *)
+  val dest_sum          = dest_Sucs_sum false
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff
+  val left_distrib      = left_add_mult_distrib RS trans
+  val prove_conv        = Int_Numeral_Base_Simprocs.prove_conv_nohyps
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [Suc_eq_add_numeral_1] @ add_ac
+  val norm_ss2 = num_ss addsimps bin_simps @ add_ac @ mult_ac
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+
+  val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps;
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq  = simplify_meta_eq
+  end;
+
+structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
+
+val combine_numerals =
+  prep_simproc ("nat_combine_numerals", ["(i::nat) + j", "Suc (i + j)"], K CombineNumerals.proc);
+
+
+(*** Applying CancelNumeralFactorFun ***)
+
+structure CancelNumeralFactorCommon =
+  struct
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff
+  val trans_tac         = fn _ => trans_tac
+
+  val norm_ss1 = num_ss addsimps
+    numeral_syms @ add_0s @ mult_1s @ [Suc_eq_add_numeral_1_left] @ add_ac
+  val norm_ss2 = num_ss addsimps bin_simps @ add_ac @ mult_ac
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1))
+    THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2))
+
+  val numeral_simp_ss = HOL_ss addsimps bin_simps
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq  = simplify_meta_eq
+  end
+
+structure DivCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
+  val cancel = nat_mult_div_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val cancel = nat_mult_eq_cancel1 RS trans
+  val neg_exchanges = false
+)
+
+structure LessCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
+  val cancel = nat_mult_less_cancel1 RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
+  val cancel = nat_mult_le_cancel1 RS trans
+  val neg_exchanges = true
+)
+
+val cancel_numeral_factors =
+  map prep_simproc
+   [("nateq_cancel_numeral_factors",
+     ["(l::nat) * m = n", "(l::nat) = m * n"],
+     K EqCancelNumeralFactor.proc),
+    ("natless_cancel_numeral_factors",
+     ["(l::nat) * m < n", "(l::nat) < m * n"],
+     K LessCancelNumeralFactor.proc),
+    ("natle_cancel_numeral_factors",
+     ["(l::nat) * m <= n", "(l::nat) <= m * n"],
+     K LeCancelNumeralFactor.proc),
+    ("natdiv_cancel_numeral_factors",
+     ["((l::nat) * m) div n", "(l::nat) div (m * n)"],
+     K DivCancelNumeralFactor.proc)];
+
+
+
+(*** Applying ExtractCommonTermFun ***)
+
+(*this version ALWAYS includes a trailing one*)
+fun long_mk_prod []        = one
+  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
+
+(*Find first term that matches u*)
+fun find_first_t past u []         = raise TERM("find_first_t", [])
+  | find_first_t past u (t::terms) =
+        if u aconv t then (rev past @ terms)
+        else find_first_t (t::past) u terms
+        handle TERM _ => find_first_t (t::past) u terms;
+
+(** Final simplification for the CancelFactor simprocs **)
+val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq  
+  [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}];
+
+fun cancel_simplify_meta_eq cancel_th ss th =
+    simplify_one ss (([th, cancel_th]) MRS trans);
+
+structure CancelFactorCommon =
+  struct
+  val mk_sum            = (fn T:typ => long_mk_prod)
+  val dest_sum          = dest_prod
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff
+  val find_first        = find_first_t []
+  val trans_tac         = fn _ => trans_tac
+  val norm_ss = HOL_ss addsimps mult_1s @ mult_ac
+  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
+  end;
+
+structure EqCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_eq_cancel_disj
+);
+
+structure LessCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
+  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_less_cancel_disj
+);
+
+structure LeCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
+  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_le_cancel_disj
+);
+
+structure DivideCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
+  val simplify_meta_eq  = cancel_simplify_meta_eq nat_mult_div_cancel_disj
+);
+
+val cancel_factor =
+  map prep_simproc
+   [("nat_eq_cancel_factor",
+     ["(l::nat) * m = n", "(l::nat) = m * n"],
+     K EqCancelFactor.proc),
+    ("nat_less_cancel_factor",
+     ["(l::nat) * m < n", "(l::nat) < m * n"],
+     K LessCancelFactor.proc),
+    ("nat_le_cancel_factor",
+     ["(l::nat) * m <= n", "(l::nat) <= m * n"],
+     K LeCancelFactor.proc),
+    ("nat_divide_cancel_factor",
+     ["((l::nat) * m) div n", "(l::nat) div (m * n)"],
+     K DivideCancelFactor.proc)];
+
+end;
+
+
+Addsimprocs Nat_Numeral_Simprocs.cancel_numerals;
+Addsimprocs [Nat_Numeral_Simprocs.combine_numerals];
+Addsimprocs Nat_Numeral_Simprocs.cancel_numeral_factors;
+Addsimprocs Nat_Numeral_Simprocs.cancel_factor;
+
+
+(*examples:
+print_depth 22;
+set timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1));
+
+(*cancel_numerals*)
+test "l +( 2) + (2) + 2 + (l + 2) + (oo  + 2) = (uu::nat)";
+test "(2*length xs < 2*length xs + j)";
+test "(2*length xs < length xs * 2 + j)";
+test "2*u = (u::nat)";
+test "2*u = Suc (u)";
+test "(i + j + 12 + (k::nat)) - 15 = y";
+test "(i + j + 12 + (k::nat)) - 5 = y";
+test "Suc u - 2 = y";
+test "Suc (Suc (Suc u)) - 2 = y";
+test "(i + j + 2 + (k::nat)) - 1 = y";
+test "(i + j + 1 + (k::nat)) - 2 = y";
+
+test "(2*x + (u*v) + y) - v*3*u = (w::nat)";
+test "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = (w::nat)";
+test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::nat)";
+test "Suc (Suc (2*x*u*v + u*4 + y)) - u = w";
+test "Suc ((u*v)*4) - v*3*u = w";
+test "Suc (Suc ((u*v)*3)) - v*3*u = w";
+
+test "(i + j + 12 + (k::nat)) = u + 15 + y";
+test "(i + j + 32 + (k::nat)) - (u + 15 + y) = zz";
+test "(i + j + 12 + (k::nat)) = u + 5 + y";
+(*Suc*)
+test "(i + j + 12 + k) = Suc (u + y)";
+test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + 41 + k)";
+test "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))";
+test "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v";
+test "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))";
+test "2*y + 3*z + 2*u = Suc (u)";
+test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)";
+test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::nat)";
+test "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)";
+test "(2*n*m) < (3*(m*n)) + (u::nat)";
+
+test "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) <= Suc 0)";
+ 
+test "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) <= length l1";
+
+test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) <= length (compT P E A ST mxr e))";
+
+test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un \<A> e) ST mxr c))))))) <= length (compT P E A ST mxr e))";
+
+
+(*negative numerals: FAIL*)
+test "(i + j + -23 + (k::nat)) < u + 15 + y";
+test "(i + j + 3 + (k::nat)) < u + -15 + y";
+test "(i + j + -12 + (k::nat)) - 15 = y";
+test "(i + j + 12 + (k::nat)) - -15 = y";
+test "(i + j + -12 + (k::nat)) - -15 = y";
+
+(*combine_numerals*)
+test "k + 3*k = (u::nat)";
+test "Suc (i + 3) = u";
+test "Suc (i + j + 3 + k) = u";
+test "k + j + 3*k + j = (u::nat)";
+test "Suc (j*i + i + k + 5 + 3*k + i*j*4) = (u::nat)";
+test "(2*n*m) + (3*(m*n)) = (u::nat)";
+(*negative numerals: FAIL*)
+test "Suc (i + j + -3 + k) = u";
+
+(*cancel_numeral_factors*)
+test "9*x = 12 * (y::nat)";
+test "(9*x) div (12 * (y::nat)) = z";
+test "9*x < 12 * (y::nat)";
+test "9*x <= 12 * (y::nat)";
+
+(*cancel_factor*)
+test "x*k = k*(y::nat)";
+test "k = k*(y::nat)";
+test "a*(b*c) = (b::nat)";
+test "a*(b*c) = d*(b::nat)*(x*a)";
+
+test "x*k < k*(y::nat)";
+test "k < k*(y::nat)";
+test "a*(b*c) < (b::nat)";
+test "a*(b*c) < d*(b::nat)*(x*a)";
+
+test "x*k <= k*(y::nat)";
+test "k <= k*(y::nat)";
+test "a*(b*c) <= (b::nat)";
+test "a*(b*c) <= d*(b::nat)*(x*a)";
+
+test "(x*k) div (k*(y::nat)) = (uu::nat)";
+test "(k) div (k*(y::nat)) = (uu::nat)";
+test "(a*(b*c)) div ((b::nat)) = (uu::nat)";
+test "(a*(b*c)) div (d*(b::nat)*(x*a)) = (uu::nat)";
+*)
+
+
+(*** Prepare linear arithmetic for nat numerals ***)
+
+local
+
+(* reduce contradictory <= to False *)
+val add_rules =
+  [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, nat_0, nat_1,
+   add_nat_number_of, diff_nat_number_of, mult_nat_number_of,
+   eq_nat_number_of, less_nat_number_of, le_number_of_eq_not_less,
+   le_Suc_number_of,le_number_of_Suc,
+   less_Suc_number_of,less_number_of_Suc,
+   Suc_eq_number_of,eq_number_of_Suc,
+   mult_Suc, mult_Suc_right,
+   eq_number_of_0, eq_0_number_of, less_0_number_of,
+   of_int_number_of_eq, of_nat_number_of_eq, nat_number_of, if_True, if_False];
+
+val simprocs = Nat_Numeral_Simprocs.combine_numerals
+  :: Nat_Numeral_Simprocs.cancel_numerals;
+
+in
+
+val nat_simprocs_setup =
+  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
+   {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
+    inj_thms = inj_thms, lessD = lessD, neqE = neqE,
+    simpset = simpset addsimps add_rules
+                      addsimprocs simprocs});
+
+end;