Version 1.0 of linear nat arithmetic.
authornipkow
Mon, 04 Jan 1999 15:07:47 +0100
changeset 6055 fdf4638bf726
parent 6054 4a4f6ad607a1
child 6056 b21813d1b701
Version 1.0 of linear nat arithmetic.
src/HOL/Arith.ML
src/HOL/Hoare/Hoare.thy
src/HOL/IsaMakefile
src/HOL/Lambda/Lambda.ML
src/HOL/Lambda/ListApplication.ML
src/HOL/List.ML
src/HOL/UNITY/Token.ML
src/HOL/arith_data.ML
--- a/src/HOL/Arith.ML	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/Arith.ML	Mon Jan 04 15:07:47 1999 +0100
@@ -626,6 +626,7 @@
 
 signature ARITH_DATA =
 sig
+  val nat_cancel_sums_add: simproc list
   val nat_cancel_sums: simproc list
   val nat_cancel_factor: simproc list
   val nat_cancel: simproc list
@@ -816,11 +817,13 @@
 val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"];
 val diff_pats = prep_pats ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"];
 
-val nat_cancel_sums = map prep_simproc
+val nat_cancel_sums_add = map prep_simproc
   [("nateq_cancel_sums", eq_pats, EqCancelSums.proc),
    ("natless_cancel_sums", less_pats, LessCancelSums.proc),
-   ("natle_cancel_sums", le_pats, LeCancelSums.proc),
-   ("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
+   ("natle_cancel_sums", le_pats, LeCancelSums.proc)];
+
+val nat_cancel_sums = nat_cancel_sums_add @
+  [prep_simproc("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
 
 val nat_cancel_factor = map prep_simproc
   [("nateq_cancel_factor", eq_pats, EqCancelFactor.proc),
@@ -853,8 +856,9 @@
 val sym = sym;
 
 val nat = Type("nat",[]);
+val boolT = Type("bool",[]);
 
-fun nnb T = T = Type("fun",[nat,Type("fun",[nat,Type("bool",[])])])
+fun nnb T = T = ([nat,nat] ---> boolT);
 
 (* Turn term into list of summand * multiplicity plus a constant *)
 fun poly(Const("Suc",_)$t, (p,i)) = poly(t, (p,i+1))
@@ -882,13 +886,14 @@
   | decomp(_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
       negate(decomp2(rel,T,lhs,rhs))
   | decomp _ = None
+
 (* reduce contradictory <= to False.
    Most of the work is done by the cancel tactics.
 *)
 val add_rules = [Zero_not_Suc,Suc_not_Zero,le_0_eq];
 
 val cancel_sums_ss = HOL_basic_ss addsimps add_rules
-                                  addsimprocs nat_cancel_sums;
+                                  addsimprocs nat_cancel_sums_add;
 
 val simp = simplify cancel_sums_ss;
 
@@ -896,129 +901,120 @@
  (fn prems => [cut_facts_tac prems 1,
                blast_tac (claset() addIs [add_le_mono]) 1]))
 ["(i <= j) & (k <= l) ==> i + k <= j + (l::nat)",
- "(i = j) & (k <= l) ==> i + k <= j + (l::nat)",
- "(i <= j) & (k = l) ==> i + k <= j + (l::nat)",
- "(i = j) & (k = l) ==> i + k <= j + (l::nat)"
+ "(i  = j) & (k <= l) ==> i + k <= j + (l::nat)",
+ "(i <= j) & (k  = l) ==> i + k <= j + (l::nat)",
+ "(i  = j) & (k  = l) ==> i + k  = j + (l::nat)"
 ];
 
+fun is_False thm =
+  let val _ $ t = #prop(rep_thm thm)
+  in t = Const("False",boolT) end;
+
+fun is_nat(t) = fastype_of1 t = nat;
+
+fun mk_nat_thm sg t =
+  let val ct = cterm_of sg t  and cn = cterm_of sg (Var(("n",0),nat))
+  in instantiate ([],[(cn,ct)]) le0 end;
+
 end;
 
 structure Fast_Nat_Arith = Fast_Lin_Arith(Nat_LA_Data);
 
 simpset_ref () := (simpset() addSolver Fast_Nat_Arith.cut_lin_arith_tac);
 
+val fast_nat_arith_tac = Fast_Nat_Arith.lin_arith_tac;
+
+(* Elimination of `-' on nat due to John Harrison *)
+Goal "P(a - b::nat) = (!d. (b = a + d --> P 0) & (a = b + d --> P d))";
+by(case_tac "a <= b" 1);
+by(Auto_tac);
+by(eres_inst_tac [("x","b-a")] allE 1);
+by(Auto_tac);
+qed "nat_diff_split";
+
+(* FIXME: K true should be replaced by a sensible test to speed things up
+   in case there are lots of irrelevant terms involved
+*)
+val nat_arith_tac =
+  refute_tac (K true) (REPEAT o split_tac[nat_diff_split])
+             ((REPEAT o etac nat_neqE) THEN' fast_nat_arith_tac);
+
 (*---------------------------------------------------------------------------*)
 (* End of proof procedures. Now go and USE them!                             *)
 (*---------------------------------------------------------------------------*)
 
-
 (*** Subtraction laws -- mostly from Clemens Ballarin ***)
 
 Goal "[| a < (b::nat); c <= a |] ==> a-c < b-c";
-by (subgoal_tac "c+(a-c) < c+(b-c)" 1);
-by (Full_simp_tac 1);
-by (subgoal_tac "c <= b" 1);
-by (blast_tac (claset() addIs [less_imp_le, le_trans]) 2);
-by (Asm_simp_tac 1);
+by(nat_arith_tac 1);
 qed "diff_less_mono";
 
 Goal "a+b < (c::nat) ==> a < c-b";
-by (dtac diff_less_mono 1);
-by (rtac le_add2 1);
-by (Asm_full_simp_tac 1);
+by(nat_arith_tac 1);
 qed "add_less_imp_less_diff";
 
 Goal "(i < j-k) = (i+k < (j::nat))";
-by (rtac iffI 1);
- by (case_tac "k <= j" 1);
-  by (dtac le_add_diff_inverse2 1);
-  by (dres_inst_tac [("k","k")] add_less_mono1 1);
-  by (Asm_full_simp_tac 1);
- by (rotate_tac 1 1);
- by (Asm_full_simp_tac 1);
-by (etac add_less_imp_less_diff 1);
+by(nat_arith_tac 1);
 qed "less_diff_conv";
 
 Goal "(j-k <= (i::nat)) = (j <= i+k)";
-by (simp_tac (simpset() addsimps [less_diff_conv, le_def]) 1);
+by(nat_arith_tac 1);
 qed "le_diff_conv";
 
 Goal "k <= j ==> (i <= j-k) = (i+k <= (j::nat))";
-by (asm_full_simp_tac
-    (simpset() delsimps [less_Suc_eq_le]
-               addsimps [less_Suc_eq_le RS sym, less_diff_conv,
-			 Suc_diff_le RS sym]) 1);
+by(nat_arith_tac 1);
 qed "le_diff_conv2";
 
 Goal "Suc i <= n ==> Suc (n - Suc i) = n - i";
-by (asm_full_simp_tac (simpset() addsimps [Suc_diff_le RS sym]) 1);
+by(nat_arith_tac 1);
 qed "Suc_diff_Suc";
 
 Goal "i <= (n::nat) ==> n - (n - i) = i";
-by (etac rev_mp 1);
-by (res_inst_tac [("m","n"),("n","i")] diff_induct 1);
-by (ALLGOALS (asm_simp_tac  (simpset() addsimps [Suc_diff_le])));
+by(nat_arith_tac 1);
 qed "diff_diff_cancel";
 Addsimps [diff_diff_cancel];
 
 Goal "k <= (n::nat) ==> m <= n + m - k";
-by (etac rev_mp 1);
-by (res_inst_tac [("m", "k"), ("n", "n")] diff_induct 1);
-by (Simp_tac 1);
-by (simp_tac (simpset() addsimps [le_add2, less_imp_le]) 1);
-by (Simp_tac 1);
+by(nat_arith_tac 1);
 qed "le_add_diff";
 
-Goal "0<k ==> j<i --> j+k-i < k";
-by (res_inst_tac [("m","j"),("n","i")] diff_induct 1);
-by (ALLGOALS Asm_simp_tac);
-qed_spec_mp "add_diff_less";
-
+Goal "[| 0<k; j<i |] ==> j+k-i < k";
+by(nat_arith_tac 1);
+qed "add_diff_less";
 
 Goal "m-1 < n ==> m <= n";
-by (exhaust_tac "m" 1);
-by (auto_tac (claset(), simpset() addsimps [Suc_le_eq]));
+by(nat_arith_tac 1);
 qed "pred_less_imp_le";
 
 Goal "j<=i ==> i - j < Suc i - j";
-by (REPEAT (etac rev_mp 1));
-by (res_inst_tac [("m","i"),("n","j")] diff_induct 1);
-by Auto_tac;
+by(nat_arith_tac 1);
 qed "diff_less_Suc_diff";
 
 Goal "i - j <= Suc i - j";
-by (res_inst_tac [("m","i"),("n","j")] diff_induct 1);
-by Auto_tac;
+by(nat_arith_tac 1);
 qed "diff_le_Suc_diff";
 AddIffs [diff_le_Suc_diff];
 
 Goal "n - Suc i <= n - i";
-by (case_tac "i<n" 1);
-by (dtac diff_Suc_less_diff 1);
-by (auto_tac (claset(), simpset() addsimps [less_imp_le, leI]));
+by(nat_arith_tac 1);
 qed "diff_Suc_le_diff";
 AddIffs [diff_Suc_le_diff];
 
 Goal "0 < n ==> (m <= n-1) = (m<n)";
-by (exhaust_tac "n" 1);
-by (auto_tac (claset(), simpset() addsimps le_simps));
+by(nat_arith_tac 1);
 qed "le_pred_eq";
 
 Goal "0 < n ==> (m-1 < n) = (m<=n)";
-by (exhaust_tac "m" 1);
-by (auto_tac (claset(), simpset() addsimps [Suc_le_eq]));
+by(nat_arith_tac 1);
 qed "less_pred_eq";
 
 (*In ordinary notation: if 0<n and n<=m then m-n < m *)
 Goal "[| 0<n; ~ m<n |] ==> m - n < m";
-by (subgoal_tac "0<n --> ~ m<n --> m - n < m" 1);
-by (Blast_tac 1);
-by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
-by (ALLGOALS(asm_simp_tac(simpset() addsimps [diff_less_Suc])));
+by(nat_arith_tac 1);
 qed "diff_less";
 
 Goal "[| 0<n; n<=m |] ==> m - n < m";
-by (asm_simp_tac (simpset() addsimps [diff_less, not_less_iff_le]) 1);
+by(nat_arith_tac 1);
 qed "le_diff_less";
 
 
@@ -1026,52 +1022,20 @@
 (** (Anti)Monotonicity of subtraction -- by Stefan Merz **)
 
 (* Monotonicity of subtraction in first argument *)
-Goal "m <= (n::nat) --> (m-l) <= (n-l)";
-by (induct_tac "n" 1);
-by (Simp_tac 1);
-by (simp_tac (simpset() addsimps [le_Suc_eq]) 1);
-by (blast_tac (claset() addIs [diff_le_Suc_diff, le_trans]) 1);
-qed_spec_mp "diff_le_mono";
+Goal "m <= (n::nat) ==> (m-l) <= (n-l)";
+by(nat_arith_tac 1);
+qed "diff_le_mono";
 
 Goal "m <= (n::nat) ==> (l-n) <= (l-m)";
-by (induct_tac "l" 1);
-by (Simp_tac 1);
-by (case_tac "n <= na" 1);
-by (subgoal_tac "m <= na" 1);
-by (asm_simp_tac (simpset() addsimps [Suc_diff_le]) 1);
-by (fast_tac (claset() addEs [le_trans]) 1);
-by (dtac not_leE 1);
-by (asm_simp_tac (simpset() addsimps [if_Suc_diff_le]) 1);
-qed_spec_mp "diff_le_mono2";
+by(nat_arith_tac 1);
+qed "diff_le_mono2";
 
 
 (*This proof requires natdiff_cancel_sums*)
-Goal "m < (n::nat) --> m<l --> (l-n) < (l-m)";
-by (induct_tac "l" 1);
-by (Simp_tac 1);
-by (Clarify_tac 1);
-by (etac less_SucE 1);
-by (Clarify_tac 2);
-by (asm_simp_tac (simpset() addsimps [Suc_le_eq]) 2);
-by (asm_simp_tac (simpset() addsimps [diff_Suc_le_Suc_diff RS le_less_trans,
-				      Suc_diff_le, less_imp_le]) 1);
-qed_spec_mp "diff_less_mono2";
-
-(** Elimination of `-' on nat due to John Harrison **)
-(*This proof requires natle_cancel_sums*)
+Goal "[| m < (n::nat); m<l |] ==> (l-n) < (l-m)";
+by(nat_arith_tac 1);
+qed "diff_less_mono2";
 
-Goal "P(a - b::nat) = (!d. (b = a + d --> P 0) & (a = b + d --> P d))";
-by(case_tac "a <= b" 1);
-by(Auto_tac);
-by(eres_inst_tac [("x","b-a")] allE 1);
-by(Auto_tac);
-qed "nat_diff_split";
-
-(*
-This is an example of the power of nat_diff_split. Many of the `-' thms in
-Arith.ML could take advantage of this, but would need to be moved.
-*)
-Goal "m-n = 0  -->  n-m = 0  -->  m=n";
-by(simp_tac (simpset() addsplits [nat_diff_split]) 1);
-qed_spec_mp "diffs0_imp_equal";
-
+Goal "[| m-n = 0; n-m = 0 |] ==>  m=n";
+by(nat_arith_tac 1);
+qed "diffs0_imp_equal";
--- a/src/HOL/Hoare/Hoare.thy	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/Hoare/Hoare.thy	Mon Jan 04 15:07:47 1999 +0100
@@ -18,12 +18,12 @@
 
 datatype
  'a com = Basic ('a fexp)         
-   | Seq ('a com) ('a com)               ("(_;/_)"      [61,60] 60)
+   | Seq ('a com) ('a com)               ("(_;/ _)"      [61,60] 60)
    | Cond ('a bexp) ('a com) ('a com)    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
    | While ('a bexp) ('a assn) ('a com)  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
   
 syntax
-  "@assign"  :: id => 'b => 'a com        ("(2_ :=/ _ )" [70,65] 61)
+  "@assign"  :: id => 'b => 'a com        ("(2_ :=/ _)" [70,65] 61)
   "@annskip" :: 'a com                    ("SKIP")
 
 translations
--- a/src/HOL/IsaMakefile	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/IsaMakefile	Mon Jan 04 15:07:47 1999 +0100
@@ -55,7 +55,7 @@
   Tools/inductive_package.ML Tools/primrec_package.ML \
   Tools/record_package.ML Tools/typedef_package.ML Trancl.ML \
   Trancl.thy Univ.ML Univ.thy Vimage.ML Vimage.thy WF.ML WF.thy \
-  WF_Rel.ML WF_Rel.thy arith_data.ML cladata.ML equalities.ML \
+  WF_Rel.ML WF_Rel.thy cladata.ML equalities.ML \
   equalities.thy hologic.ML mono.ML mono.thy simpdata.ML subset.ML \
   subset.thy thy_syntax.ML
 	@$(ISATOOL) usedir -b $(OUT)/Pure HOL
--- a/src/HOL/Lambda/Lambda.ML	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/Lambda/Lambda.ML	Mon Jan 04 15:07:47 1999 +0100
@@ -110,7 +110,6 @@
 Goal "!k. liftn (Suc n) t k = lift (liftn n t k) k";
 by (induct_tac "t" 1);
 by (ALLGOALS(asm_simp_tac(addsplit(simpset()))));
-by (blast_tac (claset() addDs [add_lessD1]) 1);
 qed "liftn_lift";
 Addsimps [liftn_lift];
 
--- a/src/HOL/Lambda/ListApplication.ML	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/Lambda/ListApplication.ML	Mon Jan 04 15:07:47 1999 +0100
@@ -88,14 +88,7 @@
 Addsimps [size_apps];
 
 Goal "[| 0 < k; m <= n |] ==> m < n+k";
-by (exhaust_tac "k" 1);
- by (Asm_full_simp_tac 1);
-by (Asm_full_simp_tac 1);
-by (rtac le_imp_less_Suc 1);
-by (exhaust_tac "n" 1);
- by (Asm_full_simp_tac 1);
-by (hyp_subst_tac 1);
-by (etac trans_le_add1 1);
+by(fast_nat_arith_tac 1);
 val lemma = result();
 
 (* a customized induction schema for $$ *)
--- a/src/HOL/List.ML	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/List.ML	Mon Jan 04 15:07:47 1999 +0100
@@ -858,7 +858,6 @@
 by (induct_tac "ns" 1);
  by (Simp_tac 1);
 by (Asm_full_simp_tac 1);
-by (blast_tac (claset() addIs [trans_le_add1]) 1);
 qed_spec_mp "start_le_sum";
 
 Goal "n : set ns ==> n <= foldl op+ 0 ns";
--- a/src/HOL/UNITY/Token.ML	Mon Dec 28 17:03:47 1998 +0100
+++ b/src/HOL/UNITY/Token.ML	Mon Jan 04 15:07:47 1999 +0100
@@ -61,8 +61,6 @@
               simpset() addsimps [diff_add_assoc2, linorder_neq_iff, 
                                   Suc_le_eq, Suc_diff_Suc, less_imp_diff_less, 
                                   add_diff_less, mod_less, mod_geq]));
-by (etac subst 1);
-by (asm_simp_tac (simpset() addsimps [add_diff_less]) 1);
 qed "nodeOrder_eq";
 
 
--- a/src/HOL/arith_data.ML	Mon Dec 28 17:03:47 1998 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,257 +0,0 @@
-(*  Title:      HOL/arith_data.ML
-    ID:         $Id$
-    Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
-
-Setup various arithmetic proof procedures.
-*)
-
-signature ARITH_DATA =
-sig
-  val nat_cancel_sums: simproc list
-  val nat_cancel_factor: simproc list
-  val nat_cancel: simproc list
-end;
-
-structure ArithData: ARITH_DATA =
-struct
-
-
-(** abstract syntax of structure nat: 0, Suc, + **)
-
-(* mk_sum, mk_norm_sum *)
-
-val one = HOLogic.mk_nat 1;
-val mk_plus = HOLogic.mk_binop "op +";
-
-fun mk_sum [] = HOLogic.zero
-  | mk_sum [t] = t
-  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
-
-(*normal form of sums: Suc (... (Suc (a + (b + ...))))*)
-fun mk_norm_sum ts =
-  let val (ones, sums) = partition (equal one) ts in
-    funpow (length ones) HOLogic.mk_Suc (mk_sum sums)
-  end;
-
-
-(* dest_sum *)
-
-val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT;
-
-fun dest_sum tm =
-  if HOLogic.is_zero tm then []
-  else
-    (case try HOLogic.dest_Suc tm of
-      Some t => one :: dest_sum t
-    | None =>
-        (case try dest_plus tm of
-          Some (t, u) => dest_sum t @ dest_sum u
-        | None => [tm]));
-
-
-(** generic proof tools **)
-
-(* prove conversions *)
-
-val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq;
-
-fun prove_conv expand_tac norm_tac sg (t, u) =
-  mk_meta_eq (prove_goalw_cterm_nocheck [] (cterm_of sg (mk_eqv (t, u)))
-    (K [expand_tac, norm_tac]))
-  handle ERROR => error ("The error(s) above occurred while trying to prove " ^
-    (string_of_cterm (cterm_of sg (mk_eqv (t, u)))));
-
-val subst_equals = prove_goal HOL.thy "[| t = s; u = t |] ==> u = s"
-  (fn prems => [cut_facts_tac prems 1, SIMPSET' asm_simp_tac 1]);
-
-
-(* rewriting *)
-
-fun simp_all rules = ALLGOALS (simp_tac (HOL_ss addsimps rules));
-
-val add_rules = [add_Suc, add_Suc_right, add_0, add_0_right];
-val mult_rules = [mult_Suc, mult_Suc_right, mult_0, mult_0_right];
-
-
-
-(** cancel common summands **)
-
-structure Sum =
-struct
-  val mk_sum = mk_norm_sum;
-  val dest_sum = dest_sum;
-  val prove_conv = prove_conv;
-  val norm_tac = simp_all add_rules THEN simp_all add_ac;
-end;
-
-fun gen_uncancel_tac rule ct =
-  rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1;
-
-
-(* nat eq *)
-
-structure EqCancelSums = CancelSumsFun
-(struct
-  open Sum;
-  val mk_bal = HOLogic.mk_eq;
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
-  val uncancel_tac = gen_uncancel_tac add_left_cancel;
-end);
-
-
-(* nat less *)
-
-structure LessCancelSums = CancelSumsFun
-(struct
-  open Sum;
-  val mk_bal = HOLogic.mk_binrel "op <";
-  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
-  val uncancel_tac = gen_uncancel_tac add_left_cancel_less;
-end);
-
-
-(* nat le *)
-
-structure LeCancelSums = CancelSumsFun
-(struct
-  open Sum;
-  val mk_bal = HOLogic.mk_binrel "op <=";
-  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
-  val uncancel_tac = gen_uncancel_tac add_left_cancel_le;
-end);
-
-
-(* nat diff *)
-
-structure DiffCancelSums = CancelSumsFun
-(struct
-  open Sum;
-  val mk_bal = HOLogic.mk_binop "op -";
-  val dest_bal = HOLogic.dest_bin "op -" HOLogic.natT;
-  val uncancel_tac = gen_uncancel_tac diff_cancel;
-end);
-
-
-
-(** cancel common factor **)
-
-structure Factor =
-struct
-  val mk_sum = mk_norm_sum;
-  val dest_sum = dest_sum;
-  val prove_conv = prove_conv;
-  val norm_tac = simp_all (add_rules @ mult_rules) THEN simp_all add_ac;
-end;
-
-fun mk_cnat n = cterm_of (sign_of Nat.thy) (HOLogic.mk_nat n);
-
-fun gen_multiply_tac rule k =
-  if k > 0 then
-    rtac (instantiate' [] [None, Some (mk_cnat (k - 1))] (rule RS subst_equals)) 1
-  else no_tac;
-
-
-(* nat eq *)
-
-structure EqCancelFactor = CancelFactorFun
-(struct
-  open Factor;
-  val mk_bal = HOLogic.mk_eq;
-  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
-  val multiply_tac = gen_multiply_tac Suc_mult_cancel1;
-end);
-
-
-(* nat less *)
-
-structure LessCancelFactor = CancelFactorFun
-(struct
-  open Factor;
-  val mk_bal = HOLogic.mk_binrel "op <";
-  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
-  val multiply_tac = gen_multiply_tac Suc_mult_less_cancel1;
-end);
-
-
-(* nat le *)
-
-structure LeCancelFactor = CancelFactorFun
-(struct
-  open Factor;
-  val mk_bal = HOLogic.mk_binrel "op <=";
-  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
-  val multiply_tac = gen_multiply_tac Suc_mult_le_cancel1;
-end);
-
-
-
-(** prepare nat_cancel simprocs **)
-
-fun prep_pat s = Thm.read_cterm (sign_of Arith.thy) (s, HOLogic.termTVar);
-val prep_pats = map prep_pat;
-
-fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
-
-val eq_pats = prep_pats ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"];
-val less_pats = prep_pats ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"];
-val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"];
-val diff_pats = prep_pats ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"];
-
-val nat_cancel_sums = map prep_simproc
-  [("nateq_cancel_sums", eq_pats, EqCancelSums.proc),
-   ("natless_cancel_sums", less_pats, LessCancelSums.proc),
-   ("natle_cancel_sums", le_pats, LeCancelSums.proc),
-   ("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
-
-val nat_cancel_factor = map prep_simproc
-  [("nateq_cancel_factor", eq_pats, EqCancelFactor.proc),
-   ("natless_cancel_factor", less_pats, LessCancelFactor.proc),
-   ("natle_cancel_factor", le_pats, LeCancelFactor.proc)];
-
-val nat_cancel = nat_cancel_factor @ nat_cancel_sums;
-
-
-end;
-
-
-open ArithData;
-
-
-context Arith.thy;
-Addsimprocs nat_cancel;
-
-
-(*This proof requires natdiff_cancel_sums*)
-Goal "m < (n::nat) --> m<l --> (l-n) < (l-m)";
-by (induct_tac "l" 1);
-by (Simp_tac 1);
-by (Clarify_tac 1);
-by (etac less_SucE 1);
-by (Clarify_tac 2);
-by (asm_simp_tac (simpset() addsimps [Suc_le_eq]) 2);
-by (asm_simp_tac (simpset() addsimps [diff_Suc_le_Suc_diff RS le_less_trans,
-				      Suc_diff_le, less_imp_le]) 1);
-qed_spec_mp "diff_less_mono2";
-
-(** Elimination of - on nat due to John Harrison **)
-(*This proof requires natle_cancel_sums*)
-
-Goal "P(a - b::nat) = (!d. (b = a + d --> P 0) & (a = b + d --> P d))";
-by(case_tac "a <= b" 1);
-by(Auto_tac);
-by(eres_inst_tac [("x","b-a")] allE 1);
-by(Auto_tac);
-qed "nat_diff_split";
-
-(*
-This is an example of the power of nat_diff_split. Many of the `-' thms in
-Arith.ML could take advantage of this, but would need to be moved.
-*)
-Goal "m-n = 0  -->  n-m = 0  -->  m=n";
-by(simp_tac (simpset() addsplits [nat_diff_split]) 1);
-qed_spec_mp "diffs0_imp_equal";
-
-use"fast_nat_decider";
-
-simpset_ref () := (simpset() addSolver
-  (fn thms => cut_facts_tac thms THEN' Fast_Nat_Decider.nat_arith_tac));