merged
authorchaieb
Fri, 08 May 2009 14:03:24 +0100
changeset 31074 710cd642650e
parent 31072 796d3d42c873 (diff)
parent 31073 4b44c4d08aa6 (current diff)
child 31075 a9d6cf6de9a8
child 31079 35b437f7ad51
merged
--- a/src/HOL/Decision_Procs/cooper_tac.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Decision_Procs/cooper_tac.ML	Fri May 08 14:03:24 2009 +0100
@@ -83,7 +83,7 @@
       addsplits [split_zdiv, split_zmod, split_div', @{thm "split_min"}, @{thm "split_max"}]
     (* Simp rules for changing (n::int) to int n *)
     val simpset1 = HOL_basic_ss
-      addsimps [nat_number_of_def, zdvd_int] @ map (fn r => r RS sym)
+      addsimps [@{thm nat_number_of_def}, zdvd_int] @ map (fn r => r RS sym)
         [@{thm int_int_eq}, @{thm zle_int}, @{thm zless_int}, @{thm zadd_int}, @{thm zmult_int}]
       addsplits [zdiff_int_split]
     (*simp rules for elimination of int n*)
--- a/src/HOL/Groebner_Basis.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Groebner_Basis.thy	Fri May 08 14:03:24 2009 +0100
@@ -635,7 +635,7 @@
 val comp_conv = (Simplifier.rewrite
 (HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"}
               addsimps ths addsimps simp_thms
-              addsimprocs field_cancel_numeral_factors
+              addsimprocs Numeral_Simprocs.field_cancel_numeral_factors
                addsimprocs [add_frac_frac_simproc, add_frac_num_simproc,
                             ord_frac_simproc]
                 addcongs [@{thm "if_weak_cong"}]))
--- a/src/HOL/Int.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Int.thy	Fri May 08 14:03:24 2009 +0100
@@ -12,10 +12,13 @@
 uses
   ("Tools/numeral.ML")
   ("Tools/numeral_syntax.ML")
+  ("Tools/int_arith.ML")
   "~~/src/Provers/Arith/assoc_fold.ML"
   "~~/src/Provers/Arith/cancel_numerals.ML"
   "~~/src/Provers/Arith/combine_numerals.ML"
-  ("Tools/int_arith.ML")
+  "~~/src/Provers/Arith/cancel_numeral_factor.ML"
+  "~~/src/Provers/Arith/extract_common_term.ML"
+  ("Tools/numeral_simprocs.ML")
 begin
 
 subsection {* The equivalence relation underlying the integers *}
@@ -1515,12 +1518,14 @@
   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
 
+use "Tools/numeral_simprocs.ML"
+
 use "Tools/int_arith.ML"
 declaration {* K Int_Arith.setup *}
 
 setup {*
   ReorientProc.add
-    (fn Const(@{const_name number_of}, _) $ _ => true | _ => false)
+    (fn Const (@{const_name number_of}, _) $ _ => true | _ => false)
 *}
 
 simproc_setup reorient_numeral ("number_of w = x") = ReorientProc.proc
--- a/src/HOL/IntDiv.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/IntDiv.thy	Fri May 08 14:03:24 2009 +0100
@@ -8,10 +8,6 @@
 
 theory IntDiv
 imports Int Divides FunDef
-uses
-  "~~/src/Provers/Arith/cancel_numeral_factor.ML"
-  "~~/src/Provers/Arith/extract_common_term.ML"
-  ("Tools/int_factor_simprocs.ML")
 begin
 
 definition divmod_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where
@@ -256,8 +252,8 @@
   val div_name = @{const_name div};
   val mod_name = @{const_name 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 mk_sum = Numeral_Simprocs.mk_sum HOLogic.intT;
+  val dest_sum = Numeral_Simprocs.dest_sum;
 
   val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}];
 
@@ -724,7 +720,6 @@
       apply (auto simp add: divmod_rel_def) 
       apply (auto simp add: algebra_simps)
       apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff)
-      apply (simp_all add: mult_less_cancel_left_disj mult_commute [of _ a])
       done
     moreover with `c \<noteq> 0` divmod_rel_div_mod have "divmod_rel b c (b div c, b mod c)" by auto
     ultimately have "divmod_rel (a * b) (a * c) (b div c, a * (b mod c))" .
@@ -1078,8 +1073,6 @@
    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)"
@@ -1334,11 +1327,6 @@
 qed
 
 
-subsection {* Simproc setup *}
-
-use "Tools/int_factor_simprocs.ML"
-
-
 subsection {* Code generation *}
 
 definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
--- a/src/HOL/IsaMakefile	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/IsaMakefile	Fri May 08 14:03:24 2009 +0100
@@ -226,19 +226,19 @@
   $(SRC)/Provers/Arith/combine_numerals.ML \
   $(SRC)/Provers/Arith/extract_common_term.ML \
   $(SRC)/Tools/Metis/metis.ML \
-  Tools/int_arith.ML \
-  Tools/int_factor_simprocs.ML \
-  Tools/nat_simprocs.ML \
   Tools/Groebner_Basis/groebner.ML \
   Tools/Groebner_Basis/misc.ML \
   Tools/Groebner_Basis/normalizer_data.ML \
   Tools/Groebner_Basis/normalizer.ML \
   Tools/atp_manager.ML \
   Tools/atp_wrapper.ML \
+  Tools/int_arith.ML \
   Tools/list_code.ML \
   Tools/meson.ML \
   Tools/metis_tools.ML \
+  Tools/nat_numeral_simprocs.ML \
   Tools/numeral.ML \
+  Tools/numeral_simprocs.ML \
   Tools/numeral_syntax.ML \
   Tools/polyhash.ML \
   Tools/Qelim/cooper_data.ML \
@@ -342,6 +342,7 @@
   Library/Random.thy Library/Quickcheck.thy	\
   Library/Poly_Deriv.thy \
   Library/Polynomial.thy \
+  Library/Preorder.thy \
   Library/Product_plus.thy \
   Library/Product_Vector.thy \
   Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML \
--- a/src/HOL/Library/Library.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Library/Library.thy	Fri May 08 14:03:24 2009 +0100
@@ -42,6 +42,7 @@
   Pocklington
   Poly_Deriv
   Polynomial
+  Preorder
   Primes
   Product_Vector
   Quickcheck
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Preorder.thy	Fri May 08 14:03:24 2009 +0100
@@ -0,0 +1,65 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Preorders with explicit equivalence relation *}
+
+theory Preorder
+imports Orderings
+begin
+
+class preorder_equiv = preorder
+begin
+
+definition equiv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
+  "equiv x y \<longleftrightarrow> x \<le> y \<and> y \<le> x"
+
+notation
+  equiv ("op ~~") and
+  equiv ("(_/ ~~ _)" [51, 51] 50)
+  
+notation (xsymbols)
+  equiv ("op \<approx>") and
+  equiv ("(_/ \<approx> _)"  [51, 51] 50)
+
+notation (HTML output)
+  equiv ("op \<approx>") and
+  equiv ("(_/ \<approx> _)"  [51, 51] 50)
+
+lemma refl [iff]:
+  "x \<approx> x"
+  unfolding equiv_def by simp
+
+lemma trans:
+  "x \<approx> y \<Longrightarrow> y \<approx> z \<Longrightarrow> x \<approx> z"
+  unfolding equiv_def by (auto intro: order_trans)
+
+lemma antisym:
+  "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x \<approx> y"
+  unfolding equiv_def ..
+
+lemma less_le: "x < y \<longleftrightarrow> x \<le> y \<and> \<not> x \<approx> y"
+  by (auto simp add: equiv_def less_le_not_le)
+
+lemma le_less: "x \<le> y \<longleftrightarrow> x < y \<or> x \<approx> y"
+  by (auto simp add: equiv_def less_le)
+
+lemma le_imp_less_or_eq: "x \<le> y \<Longrightarrow> x < y \<or> x \<approx> y"
+  by (simp add: less_le)
+
+lemma less_imp_not_eq: "x < y \<Longrightarrow> x \<approx> y \<longleftrightarrow> False"
+  by (simp add: less_le)
+
+lemma less_imp_not_eq2: "x < y \<Longrightarrow> y \<approx> x \<longleftrightarrow> False"
+  by (simp add: equiv_def less_le)
+
+lemma neq_le_trans: "\<not> a \<approx> b \<Longrightarrow> a \<le> b \<Longrightarrow> a < b"
+  by (simp add: less_le)
+
+lemma le_neq_trans: "a \<le> b \<Longrightarrow> \<not> a \<approx> b \<Longrightarrow> a < b"
+  by (simp add: less_le)
+
+lemma antisym_conv: "y \<le> x \<Longrightarrow> x \<le> y \<longleftrightarrow> x \<approx> y"
+  by (simp add: equiv_def)
+
+end
+
+end
--- a/src/HOL/Nat_Numeral.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Nat_Numeral.thy	Fri May 08 14:03:24 2009 +0100
@@ -7,7 +7,7 @@
 
 theory Nat_Numeral
 imports IntDiv
-uses ("Tools/nat_simprocs.ML")
+uses ("Tools/nat_numeral_simprocs.ML")
 begin
 
 subsection {* Numerals for natural numbers *}
@@ -455,29 +455,6 @@
 
 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*}
 
@@ -737,23 +714,6 @@
     power_number_of_odd [of "number_of v", standard]
 
 
-
-ML
-{*
-val numeral_ss = @{simpset} addsimps @{thms numerals};
-
-val nat_bin_arith_setup =
- Lin_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 @{thms neg_simps} @
-        [@{thm Suc_nat_number_of}, @{thm int_nat_number_of}]})
-*}
-
-declaration {* K 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]
@@ -912,8 +872,37 @@
 
 subsection {* Simprocs for the Naturals *}
 
-use "Tools/nat_simprocs.ML"
-declaration {* K nat_simprocs_setup *}
+use "Tools/nat_numeral_simprocs.ML"
+
+declaration {*
+let
+
+val less_eq_rules = @{thms ring_distribs} @
+  [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, @{thm nat_0}, @{thm nat_1},
+   @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of},
+   @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less},
+   @{thm le_Suc_number_of}, @{thm le_number_of_Suc},
+   @{thm less_Suc_number_of}, @{thm less_number_of_Suc},
+   @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc},
+   @{thm mult_Suc}, @{thm mult_Suc_right},
+   @{thm add_Suc}, @{thm add_Suc_right},
+   @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of},
+   @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, @{thm if_True}, @{thm if_False}];
+
+val simprocs = Nat_Numeral_Simprocs.combine_numerals :: Nat_Numeral_Simprocs.cancel_numerals;
+
+in
+
+K (Lin_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 (@{thms neg_simps} @ [@{thm Suc_nat_number_of}, @{thm int_nat_number_of}])
+      addsimps less_eq_rules
+      addsimprocs simprocs}))
+
+end
+*}
+
 
 subsubsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
 
--- a/src/HOL/Tools/int_arith.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Tools/int_arith.ML	Fri May 08 14:03:24 2009 +0100
@@ -1,420 +1,15 @@
-(* Authors: Larry Paulson and Tobias Nipkow
-
-Simprocs and decision procedure for numerals and linear arithmetic.
-*)
-
-structure Int_Numeral_Simprocs =
-struct
-
-(** 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 one_of T = Const(@{const_name HOL.one},T);
-
-(* build product with trailing 1 rather than Numeral 1 in order to avoid the
-   unnecessary restriction to type class number_ring
-   which is not required for cancellation of common factors in divisions.
-*)
-fun mk_prod T = 
-  let val one = one_of T
-  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 []        = one_of T
-  | 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 TermOrd.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) = 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, q1), (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', one_of T)) end;
-
-
-(** 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 int_ord (abs i, abs j) of
-    EQUAL => int_ord (Int.sign i, Int.sign j) 
-  | ord => ord);
-
-(*This resembles TermOrd.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 => TermOrd.typ_ord (T, U) | ord => ord)
-  | numterm_ord
-     (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) =
-     num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
-  | numterm_ord (Const(@{const_name Int.number_of}, _) $ _, _) = LESS
-  | numterm_ord (_, Const(@{const_name Int.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 TermOrd.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);
-
-val num_ss = HOL_ss settermless numtermless;
-
-(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic isn't complicated by the abstract 0 and 1.*)
-val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym];
-
-(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *)
-val add_0s =  @{thms add_0s};
-val mult_1s = @{thms mult_1s mult_1_left mult_1_right divide_1};
-
-(*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_Simprocs, such as 3 * (5 * x). *)
-val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym,
-                 @{thm add_number_of_left}, @{thm mult_number_of_left}] @
-                @{thms arith_simps} @ @{thms rel_simps};
-
-(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
-  during re-arrangement*)
-val non_add_simps =
-  subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps;
-
-(*To evaluate binary negations of coefficients*)
-val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym] @
-                   @{thms minus_bin_simps} @ @{thms pred_bin_simps};
-
-(*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 mult_minus_left}, @{thm minus_mult_right}] MRS trans |> standard;
-
-(*to extract again any uncancelled minuses*)
-val minus_from_mult_simps =
-    [@{thm minus_minus}, @{thm mult_minus_left}, @{thm mult_minus_right}];
-
-(*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];
-
-val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
-  diff_simps @ minus_simps @ @{thms add_ac}
-val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
-val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac}
+(* Author: Tobias Nipkow
 
-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         = K Arith_Data.trans_tac
-
-  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 = Arith_Data.simplify_meta_eq (add_0s @ mult_1s)
-  end;
-
-
-structure EqCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
-  val bal_add1 = @{thm eq_add_iff1} RS trans
-  val bal_add2 = @{thm eq_add_iff2} RS trans
-);
-
-structure LessCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
-  val bal_add1 = @{thm less_add_iff1} RS trans
-  val bal_add2 = @{thm less_add_iff2} RS trans
-);
-
-structure LeCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
-  val bal_add1 = @{thm le_add_iff1} RS trans
-  val bal_add2 = @{thm le_add_iff2} RS trans
-);
-
-val cancel_numerals =
-  map Arith_Data.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            = int
-  val iszero            = (fn x => x = 0)
-  val add               = op +
-  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      = @{thm combine_common_factor} RS trans
-  val prove_conv        = Arith_Data.prove_conv_nohyps
-  val trans_tac         = K Arith_Data.trans_tac
-
-  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 = Arith_Data.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            = int * int
-  val iszero            = (fn (p, 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      = @{thm combine_common_factor} RS trans
-  val prove_conv        = Arith_Data.prove_conv_nohyps
-  val trans_tac         = K Arith_Data.trans_tac
-
-  val norm_ss1a = norm_ss1 addsimps inverse_1s @ divide_simps
-  fun norm_tac ss =
-    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1a))
-    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 = Arith_Data.simplify_meta_eq (add_0s @ mult_1s @ divide_1s)
-  end;
-
-structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData);
-
-val combine_numerals =
-  Arith_Data.prep_simproc
-    ("int_combine_numerals", 
-     ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], 
-     K CombineNumerals.proc);
-
-val field_combine_numerals =
-  Arith_Data.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);
-
-(** 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 @{thms mult_ac}
-  val eq_reflection = eq_reflection
-  fun is_numeral (Const(@{const_name Int.number_of}, _) $ _) = true
-    | is_numeral _ = false;
-end;
-
-structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
-
-val assoc_fold_simproc =
-  Arith_Data.prep_simproc
-   ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
-    K Semiring_Times_Assoc.proc);
-
-end;
-
-Addsimprocs Int_Numeral_Simprocs.cancel_numerals;
-Addsimprocs [Int_Numeral_Simprocs.combine_numerals];
-Addsimprocs [Int_Numeral_Simprocs.field_combine_numerals];
-Addsimprocs [Int_Numeral_Simprocs.assoc_fold_simproc];
-
-(*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";
-*)
-
-(*** decision procedure for linear arithmetic ***)
-
-(*---------------------------------------------------------------------------*)
-(* Linear arithmetic                                                         *)
-(*---------------------------------------------------------------------------*)
-
-(*
 Instantiation of the generic linear arithmetic package for int.
 *)
 
-structure Int_Arith =
+signature INT_ARITH =
+sig
+  val fast_int_arith_simproc: simproc
+  val setup: Context.generic -> Context.generic
+end
+
+structure Int_Arith : INT_ARITH =
 struct
 
 (* Update parameters of arithmetic prover *)
@@ -491,9 +86,9 @@
 
 val nat_inj_thms = [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
 
-val int_numeral_base_simprocs = Int_Numeral_Simprocs.assoc_fold_simproc :: zero_one_idom_simproc
-  :: Int_Numeral_Simprocs.combine_numerals
-  :: Int_Numeral_Simprocs.cancel_numerals;
+val numeral_base_simprocs = Numeral_Simprocs.assoc_fold_simproc :: zero_one_idom_simproc
+  :: Numeral_Simprocs.combine_numerals
+  :: Numeral_Simprocs.cancel_numerals;
 
 val setup =
   Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
@@ -503,7 +98,7 @@
     lessD = lessD @ [@{thm zless_imp_add1_zle}],
     neqE = neqE,
     simpset = simpset addsimps add_rules
-                      addsimprocs int_numeral_base_simprocs
+                      addsimprocs numeral_base_simprocs
                       addcongs [if_weak_cong]}) #>
   arith_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) #>
   arith_discrete @{type_name Int.int}
--- a/src/HOL/Tools/int_factor_simprocs.ML	Fri May 08 14:02:33 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,391 +0,0 @@
-(*  Title:      HOL/int_factor_simprocs.ML
-    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 = [@{thm eq_number_of_eq}, @{thm less_number_of}, @{thm le_number_of}];
-
-local
-  open Int_Numeral_Simprocs
-in
-
-structure CancelNumeralFactorCommon =
-  struct
-  val mk_coeff          = mk_coeff
-  val dest_coeff        = dest_coeff 1
-  val trans_tac         = K Arith_Data.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 @{thms 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 = Arith_Data.simplify_meta_eq
-    [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left},
-      @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}];
-  end
-
-(*Version for semiring_div*)
-structure DivCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT
-  val cancel = @{thm div_mult_mult1} RS trans
-  val neg_exchanges = false
-)
-
-(*Version for fields*)
-structure DivideCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.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_mult_cancel_left} RS trans
-  val neg_exchanges = false
-)
-
-structure EqCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.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 = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
-  val cancel = @{thm mult_less_cancel_left} RS trans
-  val neg_exchanges = true
-)
-
-structure LeCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
-  val cancel = @{thm mult_le_cancel_left} RS trans
-  val neg_exchanges = true
-)
-
-val cancel_numeral_factors =
-  map Arith_Data.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::'a::{semiring_div,number_ring}) * m) div n",
-      "(l::'a::{semiring_div,number_ring}) div (m * n)"],
-     K DivCancelNumeralFactor.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 Arith_Data.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 = Arith_Data.simplify_meta_eq  
-  [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_1_eq_1}];
-
-fun cancel_simplify_meta_eq ss cancel_th th =
-    simplify_one ss (([th, cancel_th]) MRS trans);
-
-local
-  val Tp_Eq = Thm.reflexive(Thm.cterm_of (@{theory HOL}) HOLogic.Trueprop)
-  fun Eq_True_elim Eq = 
-    Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI}
-in
-fun sign_conv pos_th neg_th ss t =
-  let val T = fastype_of t;
-      val zero = Const(@{const_name HOL.zero}, T);
-      val less = Const(@{const_name HOL.less}, [T,T] ---> HOLogic.boolT);
-      val pos = less $ zero $ t and neg = less $ t $ zero
-      fun prove p =
-        Option.map Eq_True_elim (Lin_Arith.lin_arith_simproc ss p)
-        handle THM _ => NONE
-    in case prove pos of
-         SOME th => SOME(th RS pos_th)
-       | NONE => (case prove neg of
-                    SOME th => SOME(th RS neg_th)
-                  | NONE => NONE)
-    end;
-end
-
-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         = K Arith_Data.trans_tac
-  val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
-  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
-  val simplify_meta_eq  = cancel_simplify_meta_eq 
-  end;
-
-(*mult_cancel_left requires a ring with no zero divisors.*)
-structure EqCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
-  val simp_conv = K (K (SOME @{thm mult_cancel_left}))
-);
-
-(*for ordered rings*)
-structure LeCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
-  val simp_conv = sign_conv
-    @{thm mult_le_cancel_left_pos} @{thm mult_le_cancel_left_neg}
-);
-
-(*for ordered rings*)
-structure LessCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
-  val simp_conv = sign_conv
-    @{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg}
-);
-
-(*for semirings with division*)
-structure DivCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT
-  val simp_conv = K (K (SOME @{thm div_mult_mult1_if}))
-);
-
-structure ModCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binop @{const_name Divides.mod}
-  val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} HOLogic.intT
-  val simp_conv = K (K (SOME @{thm mod_mult_mult1}))
-);
-
-(*for idoms*)
-structure DvdCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
-  val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} Term.dummyT
-  val simp_conv = K (K (SOME @{thm dvd_mult_cancel_left}))
-);
-
-(*Version for all fields, including unordered ones (type complex).*)
-structure DivideCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.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 simp_conv = K (K (SOME @{thm mult_divide_mult_cancel_left_if}))
-);
-
-val cancel_factors =
-  map Arith_Data.prep_simproc
-   [("ring_eq_cancel_factor",
-     ["(l::'a::idom) * m = n",
-      "(l::'a::idom) = m * n"],
-     K EqCancelFactor.proc),
-    ("ordered_ring_le_cancel_factor",
-     ["(l::'a::ordered_ring) * m <= n",
-      "(l::'a::ordered_ring) <= m * n"],
-     K LeCancelFactor.proc),
-    ("ordered_ring_less_cancel_factor",
-     ["(l::'a::ordered_ring) * m < n",
-      "(l::'a::ordered_ring) < m * n"],
-     K LessCancelFactor.proc),
-    ("int_div_cancel_factor",
-     ["((l::'a::semiring_div) * m) div n", "(l::'a::semiring_div) div (m * n)"],
-     K DivCancelFactor.proc),
-    ("int_mod_cancel_factor",
-     ["((l::'a::semiring_div) * m) mod n", "(l::'a::semiring_div) mod (m * n)"],
-     K ModCancelFactor.proc),
-    ("dvd_cancel_factor",
-     ["((l::'a::idom) * m) dvd n", "(l::'a::idom) dvd (m * n)"],
-     K DvdCancelFactor.proc),
-    ("divide_cancel_factor",
-     ["((l::'a::{division_by_zero,field}) * m) / n",
-      "(l::'a::{division_by_zero,field}) / (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/Tools/nat_numeral_simprocs.ML	Fri May 08 14:03:24 2009 +0100
@@ -0,0 +1,538 @@
+(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Simprocs for nat numerals.
+*)
+
+signature NAT_NUMERAL_SIMPROCS =
+sig
+  val combine_numerals: simproc
+  val cancel_numerals: simproc list
+  val cancel_factors: simproc list
+  val cancel_numeral_factors: simproc list
+end;
+
+structure Nat_Numeral_Simprocs =
+struct
+
+(*Maps n to #n for n = 0, 1, 2*)
+val numeral_syms = [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, @{thm 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 = Int.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 bin_simps =
+     [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym,
+      @{thm add_nat_number_of}, @{thm nat_number_of_add_left}, 
+      @{thm diff_nat_number_of}, @{thm le_number_of_eq_not_less},
+      @{thm mult_nat_number_of}, @{thm nat_number_of_mult_left}, 
+      @{thm less_nat_number_of}, 
+      @{thm Let_number_of}, @{thm nat_number_of}] @
+     @{thms arith_simps} @ @{thms rel_simps} @ @{thms neg_simps};
+
+
+(*** 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 TermOrd.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 Int.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 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 [@{thm add_0}, @{thm 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 = [@{thm add_Suc}, @{thm add_Suc_right}, @{thm Zero_not_Suc},
+  @{thm Suc_not_Zero}, @{thm le_0_eq}];
+
+val simplify_meta_eq =
+    Arith_Data.simplify_meta_eq
+        ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right},
+          @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules);
+
+
+(*** 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         = K Arith_Data.trans_tac
+
+  val norm_ss1 = Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @
+    [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac}
+  val norm_ss2 = Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms 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 = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val bal_add1 = @{thm nat_eq_add_iff1} RS trans
+  val bal_add2 = @{thm nat_eq_add_iff2} RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
+  val bal_add1 = @{thm nat_less_add_iff1} RS trans
+  val bal_add2 = @{thm nat_less_add_iff2} RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
+  val bal_add1 = @{thm nat_le_add_iff1} RS trans
+  val bal_add2 = @{thm nat_le_add_iff2} RS trans
+);
+
+structure DiffCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Arith_Data.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 = @{thm nat_diff_add_eq1} RS trans
+  val bal_add2 = @{thm nat_diff_add_eq2} RS trans
+);
+
+
+val cancel_numerals =
+  map Arith_Data.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            = int
+  val iszero            = (fn x => x = 0)
+  val add               = op +
+  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      = @{thm left_add_mult_distrib} RS trans
+  val prove_conv        = Arith_Data.prove_conv_nohyps
+  val trans_tac         = K Arith_Data.trans_tac
+
+  val norm_ss1 = Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1}] @ @{thms add_ac}
+  val norm_ss2 = Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms 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 =
+  Arith_Data.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         = K Arith_Data.trans_tac
+
+  val norm_ss1 = Numeral_Simprocs.num_ss addsimps
+    numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac}
+  val norm_ss2 = Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms 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 = Arith_Data.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 = @{thm nat_mult_div_cancel1} RS trans
+  val neg_exchanges = false
+)
+
+structure DvdCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
+  val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
+  val cancel = @{thm nat_mult_dvd_cancel1} RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val cancel = @{thm nat_mult_eq_cancel1} RS trans
+  val neg_exchanges = false
+)
+
+structure LessCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
+  val cancel = @{thm nat_mult_less_cancel1} RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
+  val cancel = @{thm nat_mult_le_cancel1} RS trans
+  val neg_exchanges = true
+)
+
+val cancel_numeral_factors =
+  map Arith_Data.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),
+    ("natdvd_cancel_numeral_factors",
+     ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"],
+     K DvdCancelNumeralFactor.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 = Arith_Data.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 ss cancel_th 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         = K Arith_Data.trans_tac
+  val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
+  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
+  val simplify_meta_eq  = cancel_simplify_meta_eq
+  end;
+
+structure EqCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
+  val simp_conv = K(K (SOME @{thm nat_mult_eq_cancel_disj}))
+);
+
+structure LessCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
+  val simp_conv = K(K (SOME @{thm nat_mult_less_cancel_disj}))
+);
+
+structure LeCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
+  val simp_conv = K(K (SOME @{thm nat_mult_le_cancel_disj}))
+);
+
+structure DivideCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.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 simp_conv = K(K (SOME @{thm nat_mult_div_cancel_disj}))
+);
+
+structure DvdCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
+  val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
+  val simp_conv = K(K (SOME @{thm nat_mult_dvd_cancel_disj}))
+);
+
+val cancel_factor =
+  map Arith_Data.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),
+    ("nat_dvd_cancel_factor",
+     ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"],
+     K DvdCancelFactor.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)";
+*)
--- a/src/HOL/Tools/nat_simprocs.ML	Fri May 08 14:02:33 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,574 +0,0 @@
-(*  Title:      HOL/Tools/nat_simprocs.ML
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Simprocs for nat numerals.
-*)
-
-structure Nat_Numeral_Simprocs =
-struct
-
-(*Maps n to #n for n = 0, 1, 2*)
-val numeral_syms = [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, @{thm 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 = Int.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 bin_simps =
-     [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym,
-      @{thm add_nat_number_of}, @{thm nat_number_of_add_left}, 
-      @{thm diff_nat_number_of}, @{thm le_number_of_eq_not_less},
-      @{thm mult_nat_number_of}, @{thm nat_number_of_mult_left}, 
-      @{thm less_nat_number_of}, 
-      @{thm Let_number_of}, @{thm nat_number_of}] @
-     @{thms arith_simps} @ @{thms rel_simps} @ @{thms neg_simps};
-
-
-(*** 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 TermOrd.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 Int.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 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 [@{thm add_0}, @{thm 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 = [@{thm add_Suc}, @{thm add_Suc_right}, @{thm Zero_not_Suc},
-  @{thm Suc_not_Zero}, @{thm le_0_eq}];
-
-val simplify_meta_eq =
-    Arith_Data.simplify_meta_eq
-        ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right},
-          @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules);
-
-
-(*** 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         = K Arith_Data.trans_tac
-
-  val norm_ss1 = Int_Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @
-    [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac}
-  val norm_ss2 = Int_Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms 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 = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val bal_add1 = @{thm nat_eq_add_iff1} RS trans
-  val bal_add2 = @{thm nat_eq_add_iff2} RS trans
-);
-
-structure LessCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
-  val bal_add1 = @{thm nat_less_add_iff1} RS trans
-  val bal_add2 = @{thm nat_less_add_iff2} RS trans
-);
-
-structure LeCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
-  val bal_add1 = @{thm nat_le_add_iff1} RS trans
-  val bal_add2 = @{thm nat_le_add_iff2} RS trans
-);
-
-structure DiffCancelNumerals = CancelNumeralsFun
- (open CancelNumeralsCommon
-  val prove_conv = Arith_Data.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 = @{thm nat_diff_add_eq1} RS trans
-  val bal_add2 = @{thm nat_diff_add_eq2} RS trans
-);
-
-
-val cancel_numerals =
-  map Arith_Data.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            = int
-  val iszero            = (fn x => x = 0)
-  val add               = op +
-  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      = @{thm left_add_mult_distrib} RS trans
-  val prove_conv        = Arith_Data.prove_conv_nohyps
-  val trans_tac         = K Arith_Data.trans_tac
-
-  val norm_ss1 = Int_Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1}] @ @{thms add_ac}
-  val norm_ss2 = Int_Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms 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 =
-  Arith_Data.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         = K Arith_Data.trans_tac
-
-  val norm_ss1 = Int_Numeral_Simprocs.num_ss addsimps
-    numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac}
-  val norm_ss2 = Int_Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms 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 = Arith_Data.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 = @{thm nat_mult_div_cancel1} RS trans
-  val neg_exchanges = false
-)
-
-structure DvdCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
-  val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
-  val cancel = @{thm nat_mult_dvd_cancel1} RS trans
-  val neg_exchanges = false
-)
-
-structure EqCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val cancel = @{thm nat_mult_eq_cancel1} RS trans
-  val neg_exchanges = false
-)
-
-structure LessCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
-  val cancel = @{thm nat_mult_less_cancel1} RS trans
-  val neg_exchanges = true
-)
-
-structure LeCancelNumeralFactor = CancelNumeralFactorFun
- (open CancelNumeralFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
-  val cancel = @{thm nat_mult_le_cancel1} RS trans
-  val neg_exchanges = true
-)
-
-val cancel_numeral_factors =
-  map Arith_Data.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),
-    ("natdvd_cancel_numeral_factors",
-     ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"],
-     K DvdCancelNumeralFactor.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 = Arith_Data.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 ss cancel_th 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         = K Arith_Data.trans_tac
-  val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
-  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
-  val simplify_meta_eq  = cancel_simplify_meta_eq
-  end;
-
-structure EqCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_eq
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT
-  val simp_conv = K(K (SOME @{thm nat_mult_eq_cancel_disj}))
-);
-
-structure LessCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
-  val simp_conv = K(K (SOME @{thm nat_mult_less_cancel_disj}))
-);
-
-structure LeCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
-  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
-  val simp_conv = K(K (SOME @{thm nat_mult_le_cancel_disj}))
-);
-
-structure DivideCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.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 simp_conv = K(K (SOME @{thm nat_mult_div_cancel_disj}))
-);
-
-structure DvdCancelFactor = ExtractCommonTermFun
- (open CancelFactorCommon
-  val prove_conv = Arith_Data.prove_conv
-  val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
-  val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
-  val simp_conv = K(K (SOME @{thm nat_mult_dvd_cancel_disj}))
-);
-
-val cancel_factor =
-  map Arith_Data.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),
-    ("nat_dvd_cancel_factor",
-     ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"],
-     K DvdCancelFactor.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 = @{thms ring_distribs} @
-  [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, @{thm nat_0}, @{thm nat_1},
-   @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of},
-   @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less},
-   @{thm le_Suc_number_of}, @{thm le_number_of_Suc},
-   @{thm less_Suc_number_of}, @{thm less_number_of_Suc},
-   @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc},
-   @{thm mult_Suc}, @{thm mult_Suc_right},
-   @{thm add_Suc}, @{thm add_Suc_right},
-   @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of},
-   @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, @{thm if_True}, @{thm if_False}];
-
-(* Products are multiplied out during proof (re)construction via
-ring_distribs. Ideally they should remain atomic. But that is
-currently not possible because 1 is replaced by Suc 0, and then some
-simprocs start to mess around with products like (n+1)*m. The rule
-1 == Suc 0 is necessary for early parts of HOL where numerals and
-simprocs are not yet available. But then it is difficult to remove
-that rule later on, because it may find its way back in when theories
-(and thus lin-arith simpsets) are merged. Otherwise one could turn the
-rule around (Suc n = n+1) and see if that helps products being left
-alone. *)
-
-val simprocs = Nat_Numeral_Simprocs.combine_numerals
-  :: Nat_Numeral_Simprocs.cancel_numerals;
-
-in
-
-val nat_simprocs_setup =
-  Lin_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;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/numeral_simprocs.ML	Fri May 08 14:03:24 2009 +0100
@@ -0,0 +1,786 @@
+(* Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+   Copyright   2000  University of Cambridge
+
+Simprocs for the integer numerals.
+*)
+
+(*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.
+*)
+
+signature NUMERAL_SIMPROCS =
+sig
+  val mk_sum: typ -> term list -> term
+  val dest_sum: term -> term list
+
+  val assoc_fold_simproc: simproc
+  val combine_numerals: simproc
+  val cancel_numerals: simproc list
+  val cancel_factors: simproc list
+  val cancel_numeral_factors: simproc list
+  val field_combine_numerals: simproc
+  val field_cancel_numeral_factors: simproc list
+  val num_ss: simpset
+end;
+
+structure Numeral_Simprocs : NUMERAL_SIMPROCS =
+struct
+
+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 one_of T = Const(@{const_name HOL.one},T);
+
+(* build product with trailing 1 rather than Numeral 1 in order to avoid the
+   unnecessary restriction to type class number_ring
+   which is not required for cancellation of common factors in divisions.
+*)
+fun mk_prod T = 
+  let val one = one_of T
+  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 []        = one_of T
+  | 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 TermOrd.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) = 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, q1), (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', one_of T)) end;
+
+
+(** 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 int_ord (abs i, abs j) of
+    EQUAL => int_ord (Int.sign i, Int.sign j) 
+  | ord => ord);
+
+(*This resembles TermOrd.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 => TermOrd.typ_ord (T, U) | ord => ord)
+  | numterm_ord
+     (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) =
+     num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w)
+  | numterm_ord (Const(@{const_name Int.number_of}, _) $ _, _) = LESS
+  | numterm_ord (_, Const(@{const_name Int.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 TermOrd.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);
+
+val num_ss = HOL_ss settermless numtermless;
+
+(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic isn't complicated by the abstract 0 and 1.*)
+val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym];
+
+(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *)
+val add_0s =  @{thms add_0s};
+val mult_1s = @{thms mult_1s mult_1_left mult_1_right divide_1};
+
+(*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 Numeral_Simprocs, such as 3 * (5 * x). *)
+val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym,
+                 @{thm add_number_of_left}, @{thm mult_number_of_left}] @
+                @{thms arith_simps} @ @{thms rel_simps};
+
+(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
+  during re-arrangement*)
+val non_add_simps =
+  subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps;
+
+(*To evaluate binary negations of coefficients*)
+val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym] @
+                   @{thms minus_bin_simps} @ @{thms pred_bin_simps};
+
+(*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 mult_minus_left}, @{thm minus_mult_right}] MRS trans |> standard;
+
+(*to extract again any uncancelled minuses*)
+val minus_from_mult_simps =
+    [@{thm minus_minus}, @{thm mult_minus_left}, @{thm mult_minus_right}];
+
+(*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];
+
+val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @
+  diff_simps @ minus_simps @ @{thms add_ac}
+val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps
+val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac}
+
+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         = K Arith_Data.trans_tac
+
+  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 = Arith_Data.simplify_meta_eq (add_0s @ mult_1s)
+  end;
+
+
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val bal_add1 = @{thm eq_add_iff1} RS trans
+  val bal_add2 = @{thm eq_add_iff2} RS trans
+);
+
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
+  val bal_add1 = @{thm less_add_iff1} RS trans
+  val bal_add2 = @{thm less_add_iff2} RS trans
+);
+
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
+  val bal_add1 = @{thm le_add_iff1} RS trans
+  val bal_add2 = @{thm le_add_iff2} RS trans
+);
+
+val cancel_numerals =
+  map Arith_Data.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            = int
+  val iszero            = (fn x => x = 0)
+  val add               = op +
+  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      = @{thm combine_common_factor} RS trans
+  val prove_conv        = Arith_Data.prove_conv_nohyps
+  val trans_tac         = K Arith_Data.trans_tac
+
+  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 = Arith_Data.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            = int * int
+  val iszero            = (fn (p, 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      = @{thm combine_common_factor} RS trans
+  val prove_conv        = Arith_Data.prove_conv_nohyps
+  val trans_tac         = K Arith_Data.trans_tac
+
+  val norm_ss1a = norm_ss1 addsimps inverse_1s @ divide_simps
+  fun norm_tac ss =
+    ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1a))
+    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 = Arith_Data.simplify_meta_eq (add_0s @ mult_1s @ divide_1s)
+  end;
+
+structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData);
+
+val combine_numerals =
+  Arith_Data.prep_simproc
+    ("int_combine_numerals", 
+     ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], 
+     K CombineNumerals.proc);
+
+val field_combine_numerals =
+  Arith_Data.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);
+
+(** 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 @{thms mult_ac}
+  val eq_reflection = eq_reflection
+  fun is_numeral (Const(@{const_name Int.number_of}, _) $ _) = true
+    | is_numeral _ = false;
+end;
+
+structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
+
+val assoc_fold_simproc =
+  Arith_Data.prep_simproc
+   ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
+    K Semiring_Times_Assoc.proc);
+
+structure CancelNumeralFactorCommon =
+  struct
+  val mk_coeff          = mk_coeff
+  val dest_coeff        = dest_coeff 1
+  val trans_tac         = K Arith_Data.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 @{thms 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
+    [@{thm eq_number_of_eq}, @{thm less_number_of}, @{thm le_number_of}] @ simps
+  fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss))
+  val simplify_meta_eq = Arith_Data.simplify_meta_eq
+    [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left},
+      @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}];
+  end
+
+(*Version for semiring_div*)
+structure DivCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT
+  val cancel = @{thm div_mult_mult1} RS trans
+  val neg_exchanges = false
+)
+
+(*Version for fields*)
+structure DivideCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.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_mult_cancel_left} RS trans
+  val neg_exchanges = false
+)
+
+structure EqCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.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 = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
+  val cancel = @{thm mult_less_cancel_left} RS trans
+  val neg_exchanges = true
+)
+
+structure LeCancelNumeralFactor = CancelNumeralFactorFun
+ (open CancelNumeralFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
+  val cancel = @{thm mult_le_cancel_left} RS trans
+  val neg_exchanges = true
+)
+
+val cancel_numeral_factors =
+  map Arith_Data.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::'a::{semiring_div,number_ring}) * m) div n",
+      "(l::'a::{semiring_div,number_ring}) div (m * n)"],
+     K DivCancelNumeralFactor.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)];
+
+val field_cancel_numeral_factors =
+  map Arith_Data.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)]
+
+
+(** Declarations for ExtractCommonTerm **)
+
+(*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 = Arith_Data.simplify_meta_eq  
+  [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_1_eq_1}];
+
+fun cancel_simplify_meta_eq ss cancel_th th =
+    simplify_one ss (([th, cancel_th]) MRS trans);
+
+local
+  val Tp_Eq = Thm.reflexive (Thm.cterm_of @{theory HOL} HOLogic.Trueprop)
+  fun Eq_True_elim Eq = 
+    Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI}
+in
+fun sign_conv pos_th neg_th ss t =
+  let val T = fastype_of t;
+      val zero = Const(@{const_name HOL.zero}, T);
+      val less = Const(@{const_name HOL.less}, [T,T] ---> HOLogic.boolT);
+      val pos = less $ zero $ t and neg = less $ t $ zero
+      fun prove p =
+        Option.map Eq_True_elim (Lin_Arith.lin_arith_simproc ss p)
+        handle THM _ => NONE
+    in case prove pos of
+         SOME th => SOME(th RS pos_th)
+       | NONE => (case prove neg of
+                    SOME th => SOME(th RS neg_th)
+                  | NONE => NONE)
+    end;
+end
+
+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         = K Arith_Data.trans_tac
+  val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
+  fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss))
+  val simplify_meta_eq  = cancel_simplify_meta_eq 
+  end;
+
+(*mult_cancel_left requires a ring with no zero divisors.*)
+structure EqCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_eq
+  val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
+  val simp_conv = K (K (SOME @{thm mult_cancel_left}))
+);
+
+(*for ordered rings*)
+structure LeCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
+  val simp_conv = sign_conv
+    @{thm mult_le_cancel_left_pos} @{thm mult_le_cancel_left_neg}
+);
+
+(*for ordered rings*)
+structure LessCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
+  val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
+  val simp_conv = sign_conv
+    @{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg}
+);
+
+(*for semirings with division*)
+structure DivCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT
+  val simp_conv = K (K (SOME @{thm div_mult_mult1_if}))
+);
+
+structure ModCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binop @{const_name Divides.mod}
+  val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} Term.dummyT
+  val simp_conv = K (K (SOME @{thm mod_mult_mult1}))
+);
+
+(*for idoms*)
+structure DvdCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.prove_conv
+  val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
+  val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} Term.dummyT
+  val simp_conv = K (K (SOME @{thm dvd_mult_cancel_left}))
+);
+
+(*Version for all fields, including unordered ones (type complex).*)
+structure DivideCancelFactor = ExtractCommonTermFun
+ (open CancelFactorCommon
+  val prove_conv = Arith_Data.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 simp_conv = K (K (SOME @{thm mult_divide_mult_cancel_left_if}))
+);
+
+val cancel_factors =
+  map Arith_Data.prep_simproc
+   [("ring_eq_cancel_factor",
+     ["(l::'a::idom) * m = n",
+      "(l::'a::idom) = m * n"],
+     K EqCancelFactor.proc),
+    ("ordered_ring_le_cancel_factor",
+     ["(l::'a::ordered_ring) * m <= n",
+      "(l::'a::ordered_ring) <= m * n"],
+     K LeCancelFactor.proc),
+    ("ordered_ring_less_cancel_factor",
+     ["(l::'a::ordered_ring) * m < n",
+      "(l::'a::ordered_ring) < m * n"],
+     K LessCancelFactor.proc),
+    ("int_div_cancel_factor",
+     ["((l::'a::semiring_div) * m) div n", "(l::'a::semiring_div) div (m * n)"],
+     K DivCancelFactor.proc),
+    ("int_mod_cancel_factor",
+     ["((l::'a::semiring_div) * m) mod n", "(l::'a::semiring_div) mod (m * n)"],
+     K ModCancelFactor.proc),
+    ("dvd_cancel_factor",
+     ["((l::'a::idom) * m) dvd n", "(l::'a::idom) dvd (m * n)"],
+     K DvdCancelFactor.proc),
+    ("divide_cancel_factor",
+     ["((l::'a::{division_by_zero,field}) * m) / n",
+      "(l::'a::{division_by_zero,field}) / (m * n)"],
+     K DivideCancelFactor.proc)];
+
+end;
+
+Addsimprocs Numeral_Simprocs.cancel_numerals;
+Addsimprocs [Numeral_Simprocs.combine_numerals];
+Addsimprocs [Numeral_Simprocs.field_combine_numerals];
+Addsimprocs [Numeral_Simprocs.assoc_fold_simproc];
+
+(*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";
+*)
+
+Addsimprocs Numeral_Simprocs.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)";
+*)
+
+Addsimprocs Numeral_Simprocs.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/Tools/rat_arith.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Tools/rat_arith.ML	Fri May 08 14:03:24 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Real/rat_arith.ML
-    ID:         $Id$
     Author:     Lawrence C Paulson
     Copyright   2004 University of Cambridge
 
@@ -10,8 +9,6 @@
 
 local
 
-val simprocs = field_cancel_numeral_factors
-
 val simps =
  [@{thm order_less_irrefl}, @{thm neg_less_iff_less}, @{thm True_implies_equals},
   read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib},
@@ -42,7 +39,7 @@
     lessD = lessD,  (*Can't change lessD: the rats are dense!*)
     neqE =  neqE,
     simpset = simpset addsimps simps
-                      addsimprocs simprocs}) #>
+                      addsimprocs Numeral_Simprocs.field_cancel_numeral_factors}) #>
   arith_inj_const (@{const_name of_nat}, @{typ "nat => rat"}) #>
   arith_inj_const (@{const_name of_int}, @{typ "int => rat"})
 
--- a/src/HOL/Word/WordArith.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/Word/WordArith.thy	Fri May 08 14:03:24 2009 +0100
@@ -701,7 +701,8 @@
   apply (erule (2) udvd_decr0)
   done
 
-ML{*Delsimprocs cancel_factors*}
+ML {* Delsimprocs Numeral_Simprocs.cancel_factors *}
+
 lemma udvd_incr2_K: 
   "p < a + s ==> a <= a + s ==> K udvd s ==> K udvd p - a ==> a <= p ==> 
     0 < K ==> p <= p + K & p + K <= a + s"
@@ -717,7 +718,8 @@
    apply arith
   apply simp
   done
-ML{*Delsimprocs cancel_factors*}
+
+ML {* Addsimprocs Numeral_Simprocs.cancel_factors *}
 
 (* links with rbl operations *)
 lemma word_succ_rbl:
--- a/src/HOL/ex/Arith_Examples.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/ex/Arith_Examples.thy	Fri May 08 14:03:24 2009 +0100
@@ -4,7 +4,9 @@
 
 header {* Arithmetic *}
 
-theory Arith_Examples imports Main begin
+theory Arith_Examples
+imports Main
+begin
 
 text {*
   The @{text arith} method is used frequently throughout the Isabelle
@@ -35,87 +37,87 @@
            @{term Divides.div} *}
 
 lemma "(i::nat) <= max i j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::int) <= max i j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "min i j <= (i::nat)"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "min i j <= (i::int)"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "min (i::nat) j <= max i j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "min (i::int) j <= max i j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "min (i::nat) j + max i j = i + j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "min (i::int) j + max i j = i + j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::nat) < j ==> min i j < max i j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::int) < j ==> min i j < max i j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(0::int) <= abs i"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::int) <= abs i"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "abs (abs (i::int)) = abs i"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 text {* Also testing subgoals with bound variables. *}
 
 lemma "!!x. (x::nat) <= y ==> x - y = 0"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "!!x. (x::nat) - y = 0 ==> x <= y"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "!!x. ((x::nat) <= y) = (x - y = 0)"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) < y; d < 1 |] ==> x - y = d"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) < y; d < 1 |] ==> x - y - x = d - x"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(x::int) < y ==> x - y < 0"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "nat (i + j) <= nat i + nat j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "i < j ==> nat (i - j) = 0"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::nat) mod 0 = i"
   (* FIXME: need to replace 0 by its numeral representation *)
   apply (subst nat_numeral_0_eq_0 [symmetric])
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::nat) mod 1 = 0"
   (* FIXME: need to replace 1 by its numeral representation *)
   apply (subst nat_numeral_1_eq_1 [symmetric])
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::nat) mod 42 <= 41"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::int) mod 0 = i"
   (* FIXME: need to replace 0 by its numeral representation *)
   apply (subst numeral_0_eq_0 [symmetric])
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(i::int) mod 1 = 0"
   (* FIXME: need to replace 1 by its numeral representation *)
@@ -130,70 +132,70 @@
 oops
 
 lemma "-(i::int) * 1 = 0 ==> i = 0"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (0::int) < abs i; abs i * 1 < abs i * j |] ==> 1 < abs i * j"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 
 subsection {* Meta-Logic *}
 
 lemma "x < Suc y == x <= y"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "((x::nat) == z ==> x ~= y) ==> x ~= y | z ~= y"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 
 subsection {* Various Other Examples *}
 
 lemma "(x < Suc y) = (x <= y)"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) < y; y < z |] ==> x < z"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(x::nat) < y & y < z ==> x < z"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 text {* This example involves no arithmetic at all, but is solved by
   preprocessing (i.e. NNF normalization) alone. *}
 
 lemma "(P::bool) = Q ==> Q = P"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> min (x::nat) y = 0"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| P = (x = 0); (~P) = (y = 0) |] ==> max (x::nat) y = x + y"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) ~= y; a + 2 = b; a < y; y < b; a < x; x < b |] ==> False"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) > y; y > z; z > x |] ==> False"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(x::nat) - 5 > y ==> y < x"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(x::nat) ~= 0 ==> 0 < x"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) ~= y; x <= y |] ==> x < y"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (x::nat) < y; P (x - y) |] ==> P 0"
-  by (tactic {* linear_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(x - y) - (x::nat) = (x - x) - y"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "[| (a::nat) < b; c < d |] ==> (a - b) = (c - d)"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "((a::nat) - (b - (c - (d - e)))) = (a - (b - (c - (d - e))))"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
   (n = n' & n' < m) | (n = m & m < n') |
@@ -218,28 +220,28 @@
 text {* Constants. *}
 
 lemma "(0::nat) < 1"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(0::int) < 1"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(47::nat) + 11 < 08 * 15"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 lemma "(47::int) + 11 < 08 * 15"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 text {* Splitting of inequalities of different type. *}
 
 lemma "[| (a::nat) ~= b; (i::int) ~= j; a < 2; b < 2 |] ==>
   a + b <= nat (max (abs i) (abs j))"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 text {* Again, but different order. *}
 
 lemma "[| (i::int) ~= j; (a::nat) ~= b; a < 2; b < 2 |] ==>
   a + b <= nat (max (abs i) (abs j))"
-  by (tactic {* fast_arith_tac @{context} 1 *})
+  by linarith
 
 (*
 ML {* reset trace_arith; *}
--- a/src/HOL/ex/BinEx.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/ex/BinEx.thy	Fri May 08 14:03:24 2009 +0100
@@ -712,38 +712,38 @@
 by arith
 
 lemma "!!a::real. a \<le> b ==> c \<le> d ==> x + y < z ==> a + c \<le> b + d"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a < b ==> c < d ==> a - d \<le> b + (-c)"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a \<le> b ==> b + b \<le> c ==> a + a \<le> c"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a + b \<le> i + j ==> a \<le> b ==> i \<le> j ==> a + a \<le> j + j"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a + b < i + j ==> a < b ==> i < j ==> a + a < j + j"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a + b + c \<le> i + j + k \<and> a \<le> b \<and> b \<le> c \<and> i \<le> j \<and> j \<le> k --> a + a + a \<le> k + k + k"
 by arith
 
 lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
     ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a \<le> l"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
     ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a \<le> l + l + l + l"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
     ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a + a \<le> l + l + l + l + i"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
     ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a + a + a \<le> l + l + l + l + i + l"
-by (tactic "fast_arith_tac @{context} 1")
+by linarith
 
 
 subsection{*Complex Arithmetic*}
--- a/src/HOL/ex/NormalForm.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOL/ex/NormalForm.thy	Fri May 08 14:03:24 2009 +0100
@@ -10,7 +10,6 @@
 lemma "p \<longrightarrow> True" by normalization
 declare disj_assoc [code nbe]
 lemma "((P | Q) | R) = (P | (Q | R))" by normalization
-declare disj_assoc [code del]
 lemma "0 + (n::nat) = n" by normalization
 lemma "0 + Suc n = Suc n" by normalization
 lemma "Suc n + Suc m = n + Suc (Suc m)" by normalization
--- a/src/HOLCF/Pcpo.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOLCF/Pcpo.thy	Fri May 08 14:03:24 2009 +0100
@@ -13,28 +13,28 @@
 text {* The class cpo of chain complete partial orders *}
 
 class cpo = po +
-        -- {* class axiom: *}
-  assumes cpo:   "chain S \<Longrightarrow> \<exists>x :: 'a::po. range S <<| x"
+  assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x"
+begin
 
 text {* in cpo's everthing equal to THE lub has lub properties for every chain *}
 
-lemma cpo_lubI: "chain (S::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> range S <<| (\<Squnion>i. S i)"
-by (fast dest: cpo elim: lubI)
+lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)"
+  by (fast dest: cpo elim: lubI)
 
-lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = (l::'a::cpo)\<rbrakk> \<Longrightarrow> range S <<| l"
-by (blast dest: cpo intro: lubI)
+lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l"
+  by (blast dest: cpo intro: lubI)
 
 text {* Properties of the lub *}
 
-lemma is_ub_thelub: "chain (S::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)"
-by (blast dest: cpo intro: lubI [THEN is_ub_lub])
+lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)"
+  by (blast dest: cpo intro: lubI [THEN is_ub_lub])
 
 lemma is_lub_thelub:
-  "\<lbrakk>chain (S::nat \<Rightarrow> 'a::cpo); range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x"
-by (blast dest: cpo intro: lubI [THEN is_lub_lub])
+  "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x"
+  by (blast dest: cpo intro: lubI [THEN is_lub_lub])
 
 lemma lub_range_mono:
-  "\<lbrakk>range X \<subseteq> range Y; chain Y; chain (X::nat \<Rightarrow> 'a::cpo)\<rbrakk>
+  "\<lbrakk>range X \<subseteq> range Y; chain Y; chain X\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
 apply (erule is_lub_thelub)
 apply (rule ub_rangeI)
@@ -45,7 +45,7 @@
 done
 
 lemma lub_range_shift:
-  "chain (Y::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)"
+  "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)"
 apply (rule antisym_less)
 apply (rule lub_range_mono)
 apply    fast
@@ -62,7 +62,7 @@
 done
 
 lemma maxinch_is_thelub:
-  "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = ((Y i)::'a::cpo))"
+  "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)"
 apply (rule iffI)
 apply (fast intro!: thelubI lub_finch1)
 apply (unfold max_in_chain_def)
@@ -75,7 +75,7 @@
 text {* the @{text "\<sqsubseteq>"} relation between two chains is preserved by their lubs *}
 
 lemma lub_mono:
-  "\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> 
+  "\<lbrakk>chain X; chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> 
     \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
 apply (erule is_lub_thelub)
 apply (rule ub_rangeI)
@@ -87,14 +87,14 @@
 text {* the = relation between two chains is preserved by their lubs *}
 
 lemma lub_equal:
-  "\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<forall>k. X k = Y k\<rbrakk>
+  "\<lbrakk>chain X; chain Y; \<forall>k. X k = Y k\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)"
-by (simp only: expand_fun_eq [symmetric])
+  by (simp only: expand_fun_eq [symmetric])
 
 text {* more results about mono and = of lubs of chains *}
 
 lemma lub_mono2:
-  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk>
+  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain X; chain Y\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)"
 apply (erule exE)
 apply (subgoal_tac "(\<Squnion>i. X (i + Suc j)) \<sqsubseteq> (\<Squnion>i. Y (i + Suc j))")
@@ -104,12 +104,12 @@
 done
 
 lemma lub_equal2:
-  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk>
+  "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain X; chain Y\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)"
-by (blast intro: antisym_less lub_mono2 sym)
+  by (blast intro: antisym_less lub_mono2 sym)
 
 lemma lub_mono3:
-  "\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk>
+  "\<lbrakk>chain Y; chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk>
     \<Longrightarrow> (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. X i)"
 apply (erule is_lub_thelub)
 apply (rule ub_rangeI)
@@ -120,7 +120,6 @@
 done
 
 lemma ch2ch_lub:
-  fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo"
   assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
   assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
   shows "chain (\<lambda>i. \<Squnion>j. Y i j)"
@@ -130,7 +129,6 @@
 done
 
 lemma diag_lub:
-  fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo"
   assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
   assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
   shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)"
@@ -159,12 +157,12 @@
 qed
 
 lemma ex_lub:
-  fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo"
   assumes 1: "\<And>j. chain (\<lambda>i. Y i j)"
   assumes 2: "\<And>i. chain (\<lambda>j. Y i j)"
   shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)"
-by (simp add: diag_lub 1 2)
+  by (simp add: diag_lub 1 2)
 
+end
 
 subsection {* Pointed cpos *}
 
@@ -172,9 +170,9 @@
 
 class pcpo = cpo +
   assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y"
+begin
 
-definition
-  UU :: "'a::pcpo" where
+definition UU :: 'a where
   "UU = (THE x. \<forall>y. x \<sqsubseteq> y)"
 
 notation (xsymbols)
@@ -193,6 +191,8 @@
 lemma minimal [iff]: "\<bottom> \<sqsubseteq> x"
 by (rule UU_least [THEN spec])
 
+end
+
 text {* Simproc to rewrite @{term "\<bottom> = x"} to @{term "x = \<bottom>"}. *}
 
 setup {*
@@ -202,6 +202,9 @@
 
 simproc_setup reorient_bottom ("\<bottom> = x") = ReorientProc.proc
 
+context pcpo
+begin
+
 text {* useful lemmas about @{term \<bottom>} *}
 
 lemma less_UU_iff [simp]: "(x \<sqsubseteq> \<bottom>) = (x = \<bottom>)"
@@ -213,9 +216,6 @@
 lemma UU_I: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>"
 by (subst eq_UU_iff)
 
-lemma not_less2not_eq: "\<not> (x::'a::po) \<sqsubseteq> y \<Longrightarrow> x \<noteq> y"
-by auto
-
 lemma chain_UU_I: "\<lbrakk>chain Y; (\<Squnion>i. Y i) = \<bottom>\<rbrakk> \<Longrightarrow> \<forall>i. Y i = \<bottom>"
 apply (rule allI)
 apply (rule UU_I)
@@ -230,49 +230,53 @@
 done
 
 lemma chain_UU_I_inverse2: "(\<Squnion>i. Y i) \<noteq> \<bottom> \<Longrightarrow> \<exists>i::nat. Y i \<noteq> \<bottom>"
-by (blast intro: chain_UU_I_inverse)
+  by (blast intro: chain_UU_I_inverse)
 
 lemma notUU_I: "\<lbrakk>x \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> y \<noteq> \<bottom>"
-by (blast intro: UU_I)
+  by (blast intro: UU_I)
 
 lemma chain_mono2: "\<lbrakk>\<exists>j. Y j \<noteq> \<bottom>; chain Y\<rbrakk> \<Longrightarrow> \<exists>j. \<forall>i>j. Y i \<noteq> \<bottom>"
-by (blast dest: notUU_I chain_mono_less)
+  by (blast dest: notUU_I chain_mono_less)
+
+end
 
 subsection {* Chain-finite and flat cpos *}
 
 text {* further useful classes for HOLCF domains *}
 
-class finite_po = finite + po
+class chfin = po +
+  assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y"
+begin
 
-class chfin = po +
-  assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n (Y :: nat => 'a::po)"
+subclass cpo
+apply default
+apply (frule chfin)
+apply (blast intro: lub_finch1)
+done
 
-class flat = pcpo +
-  assumes ax_flat: "(x :: 'a::pcpo) \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y"
+lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y"
+  by (simp add: chfin finite_chain_def)
+
+end
 
-text {* finite partial orders are chain-finite *}
+class finite_po = finite + po
+begin
 
-instance finite_po < chfin
-apply intro_classes
+subclass chfin
+apply default
 apply (drule finite_range_imp_finch)
 apply (rule finite)
 apply (simp add: finite_chain_def)
 done
 
-text {* some properties for chfin and flat *}
-
-text {* chfin types are cpo *}
+end
 
-instance chfin < cpo
-apply intro_classes
-apply (frule chfin)
-apply (blast intro: lub_finch1)
-done
+class flat = pcpo +
+  assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y"
+begin
 
-text {* flat types are chfin *}
-
-instance flat < chfin
-apply intro_classes
+subclass chfin
+apply default
 apply (unfold max_in_chain_def)
 apply (case_tac "\<forall>i. Y i = \<bottom>")
 apply simp
@@ -283,31 +287,28 @@
 apply (blast dest: chain_mono ax_flat)
 done
 
-text {* flat subclass of chfin; @{text adm_flat} not needed *}
-
 lemma flat_less_iff:
-  fixes x y :: "'a::flat"
   shows "(x \<sqsubseteq> y) = (x = \<bottom> \<or> x = y)"
-by (safe dest!: ax_flat)
+  by (safe dest!: ax_flat)
 
-lemma flat_eq: "(a::'a::flat) \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)"
-by (safe dest!: ax_flat)
+lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)"
+  by (safe dest!: ax_flat)
 
-lemma chfin2finch: "chain (Y::nat \<Rightarrow> 'a::chfin) \<Longrightarrow> finite_chain Y"
-by (simp add: chfin finite_chain_def)
+end
 
 text {* Discrete cpos *}
 
 class discrete_cpo = sq_ord +
   assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y"
+begin
 
-subclass (in discrete_cpo) po
+subclass po
 proof qed simp_all
 
 text {* In a discrete cpo, every chain is constant *}
 
 lemma discrete_chain_const:
-  assumes S: "chain (S::nat \<Rightarrow> 'a::discrete_cpo)"
+  assumes S: "chain S"
   shows "\<exists>x. S = (\<lambda>i. x)"
 proof (intro exI ext)
   fix i :: nat
@@ -316,7 +317,7 @@
   thus "S i = S 0" by (rule sym)
 qed
 
-instance discrete_cpo < cpo
+subclass cpo
 proof
   fix S :: "nat \<Rightarrow> 'a"
   assume S: "chain S"
@@ -326,31 +327,6 @@
     by (fast intro: lub_const)
 qed
 
-text {* lemmata for improved admissibility introdution rule *}
-
-lemma infinite_chain_adm_lemma:
-  "\<lbrakk>chain Y; \<forall>i. P (Y i);  
-    \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); \<not> finite_chain Y\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk>
-      \<Longrightarrow> P (\<Squnion>i. Y i)"
-apply (case_tac "finite_chain Y")
-prefer 2 apply fast
-apply (unfold finite_chain_def)
-apply safe
-apply (erule lub_finch1 [THEN thelubI, THEN ssubst])
-apply assumption
-apply (erule spec)
-done
-
-lemma increasing_chain_adm_lemma:
-  "\<lbrakk>chain Y;  \<forall>i. P (Y i); \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i);
-    \<forall>i. \<exists>j>i. Y i \<noteq> Y j \<and> Y i \<sqsubseteq> Y j\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk>
-      \<Longrightarrow> P (\<Squnion>i. Y i)"
-apply (erule infinite_chain_adm_lemma)
-apply assumption
-apply (erule thin_rl)
-apply (unfold finite_chain_def)
-apply (unfold max_in_chain_def)
-apply (fast dest: le_imp_less_or_eq elim: chain_mono_less)
-done
+end
 
 end
--- a/src/HOLCF/Porder.thy	Fri May 08 14:02:33 2009 +0100
+++ b/src/HOLCF/Porder.thy	Fri May 08 14:03:24 2009 +0100
@@ -12,6 +12,7 @@
 
 class sq_ord =
   fixes sq_le :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+begin
 
 notation
   sq_le (infixl "<<" 55)
@@ -19,35 +20,43 @@
 notation (xsymbols)
   sq_le (infixl "\<sqsubseteq>" 55)
 
+lemma sq_ord_less_eq_trans: "\<lbrakk>a \<sqsubseteq> b; b = c\<rbrakk> \<Longrightarrow> a \<sqsubseteq> c"
+  by (rule subst)
+
+lemma sq_ord_eq_less_trans: "\<lbrakk>a = b; b \<sqsubseteq> c\<rbrakk> \<Longrightarrow> a \<sqsubseteq> c"
+  by (rule ssubst)
+
+end
+
 class po = sq_ord +
   assumes refl_less [iff]: "x \<sqsubseteq> x"
-  assumes trans_less: "\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> z\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
-  assumes antisym_less: "\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> x\<rbrakk> \<Longrightarrow> x = y"
+  assumes trans_less: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
+  assumes antisym_less: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
+begin
 
 text {* minimal fixes least element *}
 
-lemma minimal2UU[OF allI] : "\<forall>x::'a::po. uu \<sqsubseteq> x \<Longrightarrow> uu = (THE u. \<forall>y. u \<sqsubseteq> y)"
-by (blast intro: theI2 antisym_less)
+lemma minimal2UU[OF allI] : "\<forall>x. uu \<sqsubseteq> x \<Longrightarrow> uu = (THE u. \<forall>y. u \<sqsubseteq> y)"
+  by (blast intro: theI2 antisym_less)
 
 text {* the reverse law of anti-symmetry of @{term "op <<"} *}
 
-lemma antisym_less_inverse: "(x::'a::po) = y \<Longrightarrow> x \<sqsubseteq> y \<and> y \<sqsubseteq> x"
-by simp
+lemma antisym_less_inverse: "x = y \<Longrightarrow> x \<sqsubseteq> y \<and> y \<sqsubseteq> x"
+  by simp
 
-lemma box_less: "\<lbrakk>(a::'a::po) \<sqsubseteq> b; c \<sqsubseteq> a; b \<sqsubseteq> d\<rbrakk> \<Longrightarrow> c \<sqsubseteq> d"
-by (rule trans_less [OF trans_less])
-
-lemma po_eq_conv: "((x::'a::po) = y) = (x \<sqsubseteq> y \<and> y \<sqsubseteq> x)"
-by (fast elim!: antisym_less_inverse intro!: antisym_less)
+lemma box_less: "a \<sqsubseteq> b \<Longrightarrow> c \<sqsubseteq> a \<Longrightarrow> b \<sqsubseteq> d \<Longrightarrow> c \<sqsubseteq> d"
+  by (rule trans_less [OF trans_less])
 
-lemma rev_trans_less: "\<lbrakk>(y::'a::po) \<sqsubseteq> z; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
-by (rule trans_less)
+lemma po_eq_conv: "x = y \<longleftrightarrow> x \<sqsubseteq> y \<and> y \<sqsubseteq> x"
+  by (fast elim!: antisym_less_inverse intro!: antisym_less)
 
-lemma sq_ord_less_eq_trans: "\<lbrakk>a \<sqsubseteq> b; b = c\<rbrakk> \<Longrightarrow> a \<sqsubseteq> c"
-by (rule subst)
+lemma rev_trans_less: "y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z"
+  by (rule trans_less)
 
-lemma sq_ord_eq_less_trans: "\<lbrakk>a = b; b \<sqsubseteq> c\<rbrakk> \<Longrightarrow> a \<sqsubseteq> c"
-by (rule ssubst)
+lemma not_less2not_eq: "\<not> x \<sqsubseteq> y \<Longrightarrow> x \<noteq> y"
+  by auto
+
+end
 
 lemmas HOLCF_trans_rules [trans] =
   trans_less
@@ -55,49 +64,51 @@
   sq_ord_less_eq_trans
   sq_ord_eq_less_trans
 
+context po
+begin
+
 subsection {* Upper bounds *}
 
-definition
-  is_ub :: "['a set, 'a::po] \<Rightarrow> bool"  (infixl "<|" 55)  where
-  "(S <| x) = (\<forall>y. y \<in> S \<longrightarrow> y \<sqsubseteq> x)"
+definition is_ub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infixl "<|" 55) where
+  "S <| x \<longleftrightarrow> (\<forall>y. y \<in> S \<longrightarrow> y \<sqsubseteq> x)"
 
 lemma is_ubI: "(\<And>x. x \<in> S \<Longrightarrow> x \<sqsubseteq> u) \<Longrightarrow> S <| u"
-by (simp add: is_ub_def)
+  by (simp add: is_ub_def)
 
 lemma is_ubD: "\<lbrakk>S <| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u"
-by (simp add: is_ub_def)
+  by (simp add: is_ub_def)
 
 lemma ub_imageI: "(\<And>x. x \<in> S \<Longrightarrow> f x \<sqsubseteq> u) \<Longrightarrow> (\<lambda>x. f x) ` S <| u"
-unfolding is_ub_def by fast
+  unfolding is_ub_def by fast
 
 lemma ub_imageD: "\<lbrakk>f ` S <| u; x \<in> S\<rbrakk> \<Longrightarrow> f x \<sqsubseteq> u"
-unfolding is_ub_def by fast
+  unfolding is_ub_def by fast
 
 lemma ub_rangeI: "(\<And>i. S i \<sqsubseteq> x) \<Longrightarrow> range S <| x"
-unfolding is_ub_def by fast
+  unfolding is_ub_def by fast
 
 lemma ub_rangeD: "range S <| x \<Longrightarrow> S i \<sqsubseteq> x"
-unfolding is_ub_def by fast
+  unfolding is_ub_def by fast
 
 lemma is_ub_empty [simp]: "{} <| u"
-unfolding is_ub_def by fast
+  unfolding is_ub_def by fast
 
 lemma is_ub_insert [simp]: "(insert x A) <| y = (x \<sqsubseteq> y \<and> A <| y)"
-unfolding is_ub_def by fast
+  unfolding is_ub_def by fast
 
 lemma is_ub_upward: "\<lbrakk>S <| x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> S <| y"
-unfolding is_ub_def by (fast intro: trans_less)
+  unfolding is_ub_def by (fast intro: trans_less)
 
 subsection {* Least upper bounds *}
 
-definition
-  is_lub :: "['a set, 'a::po] \<Rightarrow> bool"  (infixl "<<|" 55)  where
-  "(S <<| x) = (S <| x \<and> (\<forall>u. S <| u \<longrightarrow> x \<sqsubseteq> u))"
+definition is_lub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infixl "<<|" 55) where
+  "S <<| x \<longleftrightarrow> S <| x \<and> (\<forall>u. S <| u \<longrightarrow> x \<sqsubseteq> u)"
 
-definition
-  lub :: "'a set \<Rightarrow> 'a::po" where
+definition lub :: "'a set \<Rightarrow> 'a" where
   "lub S = (THE x. S <<| x)"
 
+end
+
 syntax
   "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" ("(3LUB _:_./ _)" [0,0, 10] 10)
 
@@ -107,6 +118,9 @@
 translations
   "LUB x:A. t" == "CONST lub ((%x. t) ` A)"
 
+context po
+begin
+
 abbreviation
   Lub  (binder "LUB " 10) where
   "LUB n. t n == lub (range t)"
@@ -117,13 +131,13 @@
 text {* access to some definition as inference rule *}
 
 lemma is_lubD1: "S <<| x \<Longrightarrow> S <| x"
-unfolding is_lub_def by fast
+  unfolding is_lub_def by fast
 
 lemma is_lub_lub: "\<lbrakk>S <<| x; S <| u\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u"
-unfolding is_lub_def by fast
+  unfolding is_lub_def by fast
 
 lemma is_lubI: "\<lbrakk>S <| x; \<And>u. S <| u \<Longrightarrow> x \<sqsubseteq> u\<rbrakk> \<Longrightarrow> S <<| x"
-unfolding is_lub_def by fast
+  unfolding is_lub_def by fast
 
 text {* lubs are unique *}
 
@@ -142,54 +156,53 @@
 done
 
 lemma thelubI: "M <<| l \<Longrightarrow> lub M = l"
-by (rule unique_lub [OF lubI])
+  by (rule unique_lub [OF lubI])
 
 lemma is_lub_singleton: "{x} <<| x"
-by (simp add: is_lub_def)
+  by (simp add: is_lub_def)
 
 lemma lub_singleton [simp]: "lub {x} = x"
-by (rule thelubI [OF is_lub_singleton])
+  by (rule thelubI [OF is_lub_singleton])
 
 lemma is_lub_bin: "x \<sqsubseteq> y \<Longrightarrow> {x, y} <<| y"
-by (simp add: is_lub_def)
+  by (simp add: is_lub_def)
 
 lemma lub_bin: "x \<sqsubseteq> y \<Longrightarrow> lub {x, y} = y"
-by (rule is_lub_bin [THEN thelubI])
+  by (rule is_lub_bin [THEN thelubI])
 
 lemma is_lub_maximal: "\<lbrakk>S <| x; x \<in> S\<rbrakk> \<Longrightarrow> S <<| x"
-by (erule is_lubI, erule (1) is_ubD)
+  by (erule is_lubI, erule (1) is_ubD)
 
 lemma lub_maximal: "\<lbrakk>S <| x; x \<in> S\<rbrakk> \<Longrightarrow> lub S = x"
-by (rule is_lub_maximal [THEN thelubI])
+  by (rule is_lub_maximal [THEN thelubI])
 
 subsection {* Countable chains *}
 
-definition
+definition chain :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
   -- {* Here we use countable chains and I prefer to code them as functions! *}
-  chain :: "(nat \<Rightarrow> 'a::po) \<Rightarrow> bool" where
   "chain Y = (\<forall>i. Y i \<sqsubseteq> Y (Suc i))"
 
 lemma chainI: "(\<And>i. Y i \<sqsubseteq> Y (Suc i)) \<Longrightarrow> chain Y"
-unfolding chain_def by fast
+  unfolding chain_def by fast
 
 lemma chainE: "chain Y \<Longrightarrow> Y i \<sqsubseteq> Y (Suc i)"
-unfolding chain_def by fast
+  unfolding chain_def by fast
 
 text {* chains are monotone functions *}
 
 lemma chain_mono_less: "\<lbrakk>chain Y; i < j\<rbrakk> \<Longrightarrow> Y i \<sqsubseteq> Y j"
-by (erule less_Suc_induct, erule chainE, erule trans_less)
+  by (erule less_Suc_induct, erule chainE, erule trans_less)
 
 lemma chain_mono: "\<lbrakk>chain Y; i \<le> j\<rbrakk> \<Longrightarrow> Y i \<sqsubseteq> Y j"
-by (cases "i = j", simp, simp add: chain_mono_less)
+  by (cases "i = j", simp, simp add: chain_mono_less)
 
 lemma chain_shift: "chain Y \<Longrightarrow> chain (\<lambda>i. Y (i + j))"
-by (rule chainI, simp, erule chainE)
+  by (rule chainI, simp, erule chainE)
 
 text {* technical lemmas about (least) upper bounds of chains *}
 
 lemma is_ub_lub: "range S <<| x \<Longrightarrow> S i \<sqsubseteq> x"
-by (rule is_lubD1 [THEN ub_rangeD])
+  by (rule is_lubD1 [THEN ub_rangeD])
 
 lemma is_ub_range_shift:
   "chain S \<Longrightarrow> range (\<lambda>i. S (i + j)) <| x = range S <| x"
@@ -205,45 +218,43 @@
 
 lemma is_lub_range_shift:
   "chain S \<Longrightarrow> range (\<lambda>i. S (i + j)) <<| x = range S <<| x"
-by (simp add: is_lub_def is_ub_range_shift)
+  by (simp add: is_lub_def is_ub_range_shift)
 
 text {* the lub of a constant chain is the constant *}
 
 lemma chain_const [simp]: "chain (\<lambda>i. c)"
-by (simp add: chainI)
+  by (simp add: chainI)
 
 lemma lub_const: "range (\<lambda>x. c) <<| c"
 by (blast dest: ub_rangeD intro: is_lubI ub_rangeI)
 
 lemma thelub_const [simp]: "(\<Squnion>i. c) = c"
-by (rule lub_const [THEN thelubI])
+  by (rule lub_const [THEN thelubI])
 
 subsection {* Finite chains *}
 
-definition
+definition max_in_chain :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> bool" where
   -- {* finite chains, needed for monotony of continuous functions *}
-  max_in_chain :: "[nat, nat \<Rightarrow> 'a::po] \<Rightarrow> bool" where
-  "max_in_chain i C = (\<forall>j. i \<le> j \<longrightarrow> C i = C j)"
+  "max_in_chain i C \<longleftrightarrow> (\<forall>j. i \<le> j \<longrightarrow> C i = C j)"
 
-definition
-  finite_chain :: "(nat \<Rightarrow> 'a::po) \<Rightarrow> bool" where
+definition finite_chain :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
   "finite_chain C = (chain C \<and> (\<exists>i. max_in_chain i C))"
 
 text {* results about finite chains *}
 
 lemma max_in_chainI: "(\<And>j. i \<le> j \<Longrightarrow> Y i = Y j) \<Longrightarrow> max_in_chain i Y"
-unfolding max_in_chain_def by fast
+  unfolding max_in_chain_def by fast
 
 lemma max_in_chainD: "\<lbrakk>max_in_chain i Y; i \<le> j\<rbrakk> \<Longrightarrow> Y i = Y j"
-unfolding max_in_chain_def by fast
+  unfolding max_in_chain_def by fast
 
 lemma finite_chainI:
   "\<lbrakk>chain C; max_in_chain i C\<rbrakk> \<Longrightarrow> finite_chain C"
-unfolding finite_chain_def by fast
+  unfolding finite_chain_def by fast
 
 lemma finite_chainE:
   "\<lbrakk>finite_chain C; \<And>i. \<lbrakk>chain C; max_in_chain i C\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
-unfolding finite_chain_def by fast
+  unfolding finite_chain_def by fast
 
 lemma lub_finch1: "\<lbrakk>chain C; max_in_chain i C\<rbrakk> \<Longrightarrow> range C <<| C i"
 apply (rule is_lubI)
@@ -311,11 +322,11 @@
 done
 
 lemma bin_chain: "x \<sqsubseteq> y \<Longrightarrow> chain (\<lambda>i. if i=0 then x else y)"
-by (rule chainI, simp)
+  by (rule chainI, simp)
 
 lemma bin_chainmax:
   "x \<sqsubseteq> y \<Longrightarrow> max_in_chain (Suc 0) (\<lambda>i. if i=0 then x else y)"
-unfolding max_in_chain_def by simp
+  unfolding max_in_chain_def by simp
 
 lemma lub_bin_chain:
   "x \<sqsubseteq> y \<Longrightarrow> range (\<lambda>i::nat. if i=0 then x else y) <<| y"
@@ -328,36 +339,35 @@
 text {* the maximal element in a chain is its lub *}
 
 lemma lub_chain_maxelem: "\<lbrakk>Y i = c; \<forall>i. Y i \<sqsubseteq> c\<rbrakk> \<Longrightarrow> lub (range Y) = c"
-by (blast dest: ub_rangeD intro: thelubI is_lubI ub_rangeI)
+  by (blast dest: ub_rangeD intro: thelubI is_lubI ub_rangeI)
 
 subsection {* Directed sets *}
 
-definition
-  directed :: "'a::po set \<Rightarrow> bool" where
-  "directed S = ((\<exists>x. x \<in> S) \<and> (\<forall>x\<in>S. \<forall>y\<in>S. \<exists>z\<in>S. x \<sqsubseteq> z \<and> y \<sqsubseteq> z))"
+definition directed :: "'a set \<Rightarrow> bool" where
+  "directed S \<longleftrightarrow> (\<exists>x. x \<in> S) \<and> (\<forall>x\<in>S. \<forall>y\<in>S. \<exists>z\<in>S. x \<sqsubseteq> z \<and> y \<sqsubseteq> z)"
 
 lemma directedI:
   assumes 1: "\<exists>z. z \<in> S"
   assumes 2: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> \<exists>z\<in>S. x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
   shows "directed S"
-unfolding directed_def using prems by fast
+  unfolding directed_def using prems by fast
 
 lemma directedD1: "directed S \<Longrightarrow> \<exists>z. z \<in> S"
-unfolding directed_def by fast
+  unfolding directed_def by fast
 
 lemma directedD2: "\<lbrakk>directed S; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> \<exists>z\<in>S. x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
-unfolding directed_def by fast
+  unfolding directed_def by fast
 
 lemma directedE1:
   assumes S: "directed S"
   obtains z where "z \<in> S"
-by (insert directedD1 [OF S], fast)
+  by (insert directedD1 [OF S], fast)
 
 lemma directedE2:
   assumes S: "directed S"
   assumes x: "x \<in> S" and y: "y \<in> S"
   obtains z where "z \<in> S" "x \<sqsubseteq> z" "y \<sqsubseteq> z"
-by (insert directedD2 [OF S x y], fast)
+  by (insert directedD2 [OF S x y], fast)
 
 lemma directed_finiteI:
   assumes U: "\<And>U. \<lbrakk>finite U; U \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>z\<in>S. U <| z"
@@ -395,13 +405,13 @@
 qed
 
 lemma not_directed_empty [simp]: "\<not> directed {}"
-by (rule notI, drule directedD1, simp)
+  by (rule notI, drule directedD1, simp)
 
 lemma directed_singleton: "directed {x}"
-by (rule directedI, auto)
+  by (rule directedI, auto)
 
 lemma directed_bin: "x \<sqsubseteq> y \<Longrightarrow> directed {x, y}"
-by (rule directedI, auto)
+  by (rule directedI, auto)
 
 lemma directed_chain: "chain S \<Longrightarrow> directed (range S)"
 apply (rule directedI)
@@ -412,4 +422,33 @@
 apply simp
 done
 
+text {* lemmata for improved admissibility introdution rule *}
+
+lemma infinite_chain_adm_lemma:
+  "\<lbrakk>chain Y; \<forall>i. P (Y i);  
+    \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); \<not> finite_chain Y\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk>
+      \<Longrightarrow> P (\<Squnion>i. Y i)"
+apply (case_tac "finite_chain Y")
+prefer 2 apply fast
+apply (unfold finite_chain_def)
+apply safe
+apply (erule lub_finch1 [THEN thelubI, THEN ssubst])
+apply assumption
+apply (erule spec)
+done
+
+lemma increasing_chain_adm_lemma:
+  "\<lbrakk>chain Y;  \<forall>i. P (Y i); \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i);
+    \<forall>i. \<exists>j>i. Y i \<noteq> Y j \<and> Y i \<sqsubseteq> Y j\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk>
+      \<Longrightarrow> P (\<Squnion>i. Y i)"
+apply (erule infinite_chain_adm_lemma)
+apply assumption
+apply (erule thin_rl)
+apply (unfold finite_chain_def)
+apply (unfold max_in_chain_def)
+apply (fast dest: le_imp_less_or_eq elim: chain_mono_less)
+done
+
 end
+
+end
\ No newline at end of file
--- a/src/Tools/code/code_ml.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/Tools/code/code_ml.ML	Fri May 08 14:03:24 2009 +0100
@@ -958,10 +958,7 @@
 fun eval some_target reff postproc thy t args =
   let
     val ctxt = ProofContext.init thy;
-    val _ = if null (Term.add_frees t []) then () else error ("Term "
-      ^ quote (Syntax.string_of_term_global thy t)
-      ^ " to be evaluated contains free variables");
-    fun evaluator naming program (((_, (_, ty)), _), t) deps =
+    fun evaluator naming program ((_, (_, ty)), t) deps =
       let
         val _ = if Code_Thingol.contains_dictvar t then
           error "Term to be evaluated contains free dictionaries" else ();
--- a/src/Tools/code/code_thingol.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/Tools/code/code_thingol.ML	Fri May 08 14:03:24 2009 +0100
@@ -85,10 +85,10 @@
   val consts_program: theory -> string list -> string list * (naming * program)
   val cached_program: theory -> naming * program
   val eval_conv: theory -> (sort -> sort)
-    -> (naming -> program -> (((string * sort) list * typscheme) * (string * itype) list) * iterm -> string list -> cterm -> thm)
+    -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> cterm -> thm)
     -> cterm -> thm
   val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a)
-    -> (naming -> program -> (((string * sort) list * typscheme) * (string * itype) list) * iterm -> string list -> 'a)
+    -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)
     -> term -> 'a
 end;
 
@@ -755,7 +755,7 @@
 
 (* value evaluation *)
 
-fun ensure_value thy algbr funcgr (fs, t) =
+fun ensure_value thy algbr funcgr t =
   let
     val ty = fastype_of t;
     val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =)
@@ -766,7 +766,7 @@
       ##>> translate_term thy algbr funcgr NONE t
       #>> (fn ((vs, ty), t) => Fun
         (Term.dummy_patternN, ((vs, ty), [(([], t), (Drule.dummy_thm, true))])));
-    fun term_value fs (dep, (naming, program1)) =
+    fun term_value (dep, (naming, program1)) =
       let
         val Fun (_, (vs_ty, [(([], t), _)])) =
           Graph.get_node program1 Term.dummy_patternN;
@@ -774,22 +774,19 @@
         val program2 = Graph.del_nodes [Term.dummy_patternN] program1;
         val deps_all = Graph.all_succs program2 deps;
         val program3 = Graph.subgraph (member (op =) deps_all) program2;
-      in (((naming, program3), (((vs_ty, fs), t), deps)), (dep, (naming, program2))) end;
+      in (((naming, program3), ((vs_ty, t), deps)), (dep, (naming, program2))) end;
   in
     ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN
     #> snd
-    #> fold_map (fn (v, ty) => translate_typ thy algbr funcgr ty
-         #-> (fn ty' => pair (v, ty'))) fs
-    #-> (fn fs' => term_value fs')
+    #> term_value
   end;
 
 fun base_evaluator thy evaluator algebra funcgr vs t =
   let
-    val fs = Term.add_frees t [];
-    val (((naming, program), ((((vs', ty'), fs'), t'), deps)), _) =
-      invoke_generation thy (algebra, funcgr) ensure_value (fs, t);
+    val (((naming, program), (((vs', ty'), t'), deps)), _) =
+      invoke_generation thy (algebra, funcgr) ensure_value t;
     val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs';
-  in evaluator naming program (((vs'', (vs', ty')), fs'), t') deps end;
+  in evaluator naming program ((vs'', (vs', ty')), t') deps end;
 
 fun eval_conv thy prep_sort = Code_Wellsorted.eval_conv thy prep_sort o base_evaluator thy;
 fun eval thy prep_sort postproc = Code_Wellsorted.eval thy prep_sort postproc o base_evaluator thy;
--- a/src/Tools/code/code_wellsorted.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/Tools/code/code_wellsorted.ML	Fri May 08 14:03:24 2009 +0100
@@ -293,19 +293,22 @@
 fun obtain thy cs ts = apsnd snd
   (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts));
 
-fun prepare_sorts prep_sort (Const (c, ty)) = Const (c, map_type_tfree
-      (fn (v, sort) => TFree (v, prep_sort sort)) ty)
+fun prepare_sorts_typ prep_sort
+  = map_type_tfree (fn (v, sort) => TFree (v, prep_sort sort));
+
+fun prepare_sorts prep_sort (Const (c, ty)) =
+      Const (c, prepare_sorts_typ prep_sort ty)
   | prepare_sorts prep_sort (t1 $ t2) =
       prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2
   | prepare_sorts prep_sort (Abs (v, ty, t)) =
-      Abs (v, Type.strip_sorts ty, prepare_sorts prep_sort t)
-  | prepare_sorts _ (Term.Free (v, ty)) = Term.Free (v, Type.strip_sorts ty)
+      Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t)
   | prepare_sorts _ (t as Bound _) = t;
 
 fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct =
   let
+    val pp = Syntax.pp_global thy;
     val ct = cterm_of proto_ct;
-    val _ = (Term.map_types Type.no_tvars o Sign.no_vars (Syntax.pp_global thy))
+    val _ = (Sign.no_frees pp o map_types (K dummyT) o Sign.no_vars pp)
       (Thm.term_of ct);
     val thm = Code.preprocess_conv thy ct;
     val ct' = Thm.rhs_of thm;
@@ -313,6 +316,7 @@
     val vs = Term.add_tfrees t' [];
     val consts = fold_aterms
       (fn Const (c, _) => insert (op =) c | _ => I) t' [];
+ 
     val t'' = prepare_sorts prep_sort t';
     val (algebra', eqngr') = obtain thy consts [t''];
   in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end;
--- a/src/Tools/nbe.ML	Fri May 08 14:02:33 2009 +0100
+++ b/src/Tools/nbe.ML	Fri May 08 14:03:24 2009 +0100
@@ -11,7 +11,6 @@
 
   datatype Univ =
       Const of int * Univ list               (*named (uninterpreted) constants*)
-    | Free of string * Univ list             (*free (uninterpreted) variables*)
     | DFree of string * int                  (*free (uninterpreted) dictionary parameters*)
     | BVar of int * Univ list
     | Abs of (int * (Univ list -> Univ)) * Univ list
@@ -57,14 +56,12 @@
 
 datatype Univ =
     Const of int * Univ list           (*named (uninterpreted) constants*)
-  | Free of string * Univ list         (*free variables*)
   | DFree of string * int              (*free (uninterpreted) dictionary parameters*)
   | BVar of int * Univ list            (*bound variables, named*)
   | Abs of (int * (Univ list -> Univ)) * Univ list
                                        (*abstractions as closures*);
 
 fun same (Const (k, xs)) (Const (l, ys)) = k = l andalso sames xs ys
-  | same (Free (s, xs)) (Free (t, ys)) = s = t andalso sames xs ys
   | same (DFree (s, k)) (DFree (t, l)) = s = t andalso k = l
   | same (BVar (k, xs)) (BVar (l, ys)) = k = l andalso sames xs ys
   | same _ _ = false
@@ -80,7 +77,6 @@
         | GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*)
       end
   | apps (Const (name, xs)) ys = Const (name, ys @ xs)
-  | apps (Free (name, xs)) ys = Free (name, ys @ xs)
   | apps (BVar (n, xs)) ys = BVar (n, ys @ xs);
 
 
@@ -352,14 +348,12 @@
 
 fun eval_term ctxt gr deps (vs : (string * sort) list, t) =
   let 
-    val frees = Code_Thingol.fold_unbound_varnames (insert (op =)) t []
-    val frees' = map (fn v => Free (v, [])) frees;
     val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs;
   in
-    ("", (vs, [(map IVar frees, t)]))
+    ("", (vs, [([], t)]))
     |> singleton (compile_eqnss ctxt gr deps)
     |> snd
-    |> (fn t => apps t (rev (dict_frees @ frees')))
+    |> (fn t => apps t (rev dict_frees))
   end;
 
 (* reification *)
@@ -399,8 +393,6 @@
             val T' = map_type_tfree (fn (v, _) => TypeInfer.param typidx (v, [])) T;
             val typidx' = typidx + 1;
           in of_apps bounds (Term.Const (c, T'), ts') typidx' end
-      | of_univ bounds (Free (name, ts)) typidx =
-          of_apps bounds (Term.Free (name, dummyT), ts) typidx
       | of_univ bounds (BVar (n, ts)) typidx =
           of_apps bounds (Bound (bounds - n - 1), ts) typidx
       | of_univ bounds (t as Abs _) typidx =
@@ -440,15 +432,12 @@
 
 (* evaluation with type reconstruction *)
 
-fun normalize thy naming program (((vs0, (vs, ty)), fs), t) deps =
+fun normalize thy naming program ((vs0, (vs, ty)), t) deps =
   let
     fun subst_const f = map_aterms (fn t as Term.Const (c, ty) => Term.Const (f c, ty)
       | t => t);
     val resubst_triv_consts = subst_const (Code_Unit.resubst_alias thy);
     val ty' = typ_of_itype program vs0 ty;
-    val fs' = (map o apsnd) (typ_of_itype program vs0) fs;
-    val type_frees = Term.map_aterms (fn (t as Term.Free (s, _)) =>
-      Term.Free (s, (the o AList.lookup (op =) fs') s) | t => t);
     fun type_infer t =
       singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I
         (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE) Name.context 0)
@@ -461,8 +450,6 @@
     compile_eval thy naming program (vs, t) deps
     |> tracing (fn t => "Normalized:\n" ^ string_of_term t)
     |> resubst_triv_consts
-    |> type_frees
-    |> tracing (fn t => "Vars typed:\n" ^ string_of_term t)
     |> type_infer
     |> tracing (fn t => "Types inferred:\n" ^ string_of_term t)
     |> check_tvars
@@ -482,18 +469,41 @@
   in Thm.mk_binop eq lhs rhs end;
 
 val (_, raw_norm_oracle) = Context.>>> (Context.map_theory_result
-  (Thm.add_oracle (Binding.name "norm", fn (thy, naming, program, vsp_ty_fs_t, deps, ct) =>
-    mk_equals thy ct (normalize thy naming program vsp_ty_fs_t deps))));
+  (Thm.add_oracle (Binding.name "norm", fn (thy, naming, program, vsp_ty_t, deps, ct) =>
+    mk_equals thy ct (normalize thy naming program vsp_ty_t deps))));
+
+fun norm_oracle thy naming program vsp_ty_t deps ct =
+  raw_norm_oracle (thy, naming, program, vsp_ty_t, deps, ct);
 
-fun norm_oracle thy naming program vsp_ty_fs_t deps ct =
-  raw_norm_oracle (thy, naming, program, vsp_ty_fs_t, deps, ct);
+fun no_frees_conv conv ct =
+  let
+    val frees = Thm.add_cterm_frees ct [];
+    fun apply_beta free thm = Thm.combination thm (Thm.reflexive free)
+      |> Conv.fconv_rule (Conv.arg_conv (Conv.try_conv (Thm.beta_conversion false)))
+      |> Conv.fconv_rule (Conv.arg1_conv (Thm.beta_conversion false));
+  in
+    ct
+    |> fold_rev Thm.cabs frees
+    |> conv
+    |> fold apply_beta frees
+  end;
 
-fun norm_conv ct =
+fun no_frees_rew rew t =
+  let
+    val frees = map Free (Term.add_frees t []);
+  in
+    t
+    |> fold_rev lambda frees
+    |> rew
+    |> (fn t' => Term.betapplys (t', frees))
+  end;
+
+val norm_conv = no_frees_conv (fn ct =>
   let
     val thy = Thm.theory_of_cterm ct;
-  in Code_Thingol.eval_conv thy (add_triv_classes thy) (norm_oracle thy) ct end;
+  in Code_Thingol.eval_conv thy (add_triv_classes thy) (norm_oracle thy) ct end);
 
-fun norm thy = Code_Thingol.eval thy (add_triv_classes thy) I (normalize thy);
+fun norm thy = no_frees_rew (Code_Thingol.eval thy (add_triv_classes thy) I (normalize thy));
 
 (* evaluation command *)