src/HOL/arith_data.ML
author wenzelm
Wed, 26 Nov 1997 17:52:53 +0100
changeset 4309 c98f44948d86
parent 4295 81132a38996a
child 4310 cad4f24be60b
permissions -rw-r--r--
separate lists of simprocs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/arith_data.ML
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     4
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     5
Setup various arithmetic proof procedures.
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     6
*)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     7
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     8
signature ARITH_DATA =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
     9
sig
4309
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
    10
  val nat_cancel_sums: simproc list
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
    11
  val nat_cancel_factor: simproc list
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    12
  val nat_cancel: simproc list
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    13
end;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    14
4309
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
    15
structure ArithData: ARITH_DATA =
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    16
struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    17
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    18
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    19
(** abstract syntax of structure nat: 0, Suc, + **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    20
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    21
(* mk_sum *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    22
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    23
val mk_plus = HOLogic.mk_binop "op +";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    24
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    25
fun mk_sum [] = HOLogic.zero
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    26
  | mk_sum [t] = t
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    27
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    28
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    29
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    30
(* dest_sum *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    31
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    32
val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    33
val one = HOLogic.mk_nat 1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    34
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    35
fun dest_sum tm =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    36
  if HOLogic.is_zero tm then []
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    37
  else
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    38
    (case try HOLogic.dest_Suc tm of
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    39
      Some t => one :: dest_sum t
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    40
    | None =>
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    41
        (case try dest_plus tm of
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    42
          Some (t, u) => dest_sum t @ dest_sum u
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    43
        | None => [tm]));
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    44
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    45
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    46
(** generic proof tools **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    47
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    48
(* prove conversions *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    49
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    50
val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    51
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    52
fun prove_conv expand_tac norm_tac sg (t, u) =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    53
  mk_meta_eq (prove_goalw_cterm [] (cterm_of sg (mk_eqv (t, u)))
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    54
    (K [expand_tac, norm_tac]))
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    55
  handle ERROR => error ("The error(s) above occurred while trying to prove " ^
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    56
    (string_of_cterm (cterm_of sg (mk_eqv (t, u)))));
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    57
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    58
val subst_equals = prove_goal HOL.thy "[| t = s; u = t |] ==> u = s"
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    59
  (fn prems => [cut_facts_tac prems 1, SIMPSET' asm_simp_tac 1]);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    60
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    61
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    62
(* rewriting *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    63
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    64
fun simp_all rules = ALLGOALS (simp_tac (HOL_ss addsimps rules));
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    65
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    66
val add_rules = [add_Suc, add_Suc_right, add_0, add_0_right];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    67
val mult_rules = [mult_Suc, mult_Suc_right, mult_0, mult_0_right];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    68
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    69
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    70
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    71
(** cancel common summands **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    72
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    73
structure Sum =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    74
struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    75
  val mk_sum = mk_sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    76
  val dest_sum = dest_sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    77
  val prove_conv = prove_conv;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    78
  val norm_tac = simp_all add_rules THEN simp_all add_ac;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    79
end;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    80
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    81
fun gen_uncancel_tac rule ct =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    82
  rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    83
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    84
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    85
(* nat eq *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    86
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    87
structure EqCancelSums = CancelSumsFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    88
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    89
  open Sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    90
  val mk_bal = HOLogic.mk_eq;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    91
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    92
  val uncancel_tac = gen_uncancel_tac add_left_cancel;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    93
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    94
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    95
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    96
(* nat less *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    97
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    98
structure LessCancelSums = CancelSumsFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    99
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   100
  open Sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   101
  val mk_bal = HOLogic.mk_binrel "op <";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   102
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   103
  val uncancel_tac = gen_uncancel_tac add_left_cancel_less;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   104
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   105
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   106
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   107
(* nat le *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   108
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   109
structure LeCancelSums = CancelSumsFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   110
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   111
  open Sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   112
  val mk_bal = HOLogic.mk_binrel "op <=";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   113
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   114
  val uncancel_tac = gen_uncancel_tac add_left_cancel_le;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   115
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   116
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   117
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   118
(* nat diff *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   119
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   120
(* FIXME *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   121
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   122
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   123
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   124
(** cancel common factor **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   125
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   126
structure Factor =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   127
struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   128
  val mk_sum = mk_sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   129
  val dest_sum = dest_sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   130
  val prove_conv = prove_conv;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   131
  val norm_tac = simp_all (add_rules @ mult_rules) THEN simp_all add_ac;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   132
end;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   133
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   134
fun mk_cnat n = cterm_of (sign_of Nat.thy) (HOLogic.mk_nat n);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   135
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   136
fun gen_multiply_tac rule k =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   137
  if k > 0 then
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   138
    rtac (instantiate' [] [None, Some (mk_cnat (k - 1))] (rule RS subst_equals)) 1
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   139
  else no_tac;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   140
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   141
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   142
(* nat eq *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   143
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   144
structure EqCancelFactor = CancelFactorFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   145
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   146
  open Factor;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   147
  val mk_bal = HOLogic.mk_eq;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   148
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   149
  val multiply_tac = gen_multiply_tac Suc_mult_cancel1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   150
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   151
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   152
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   153
(* nat less *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   154
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   155
structure LessCancelFactor = CancelFactorFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   156
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   157
  open Factor;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   158
  val mk_bal = HOLogic.mk_binrel "op <";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   159
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   160
  val multiply_tac = gen_multiply_tac Suc_mult_less_cancel1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   161
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   162
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   163
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   164
(* nat le *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   165
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   166
structure LeCancelFactor = CancelFactorFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   167
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   168
  open Factor;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   169
  val mk_bal = HOLogic.mk_binrel "op <=";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   170
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   171
  val multiply_tac = gen_multiply_tac Suc_mult_le_cancel1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   172
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   173
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   174
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   175
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   176
(** prepare nat_cancel simprocs **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   177
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   178
fun prep_pat s = Thm.read_cterm (sign_of Arith.thy) (s, HOLogic.termTVar);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   179
val prep_pats = map prep_pat;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   180
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   181
fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   182
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   183
val eq_pats = prep_pats ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   184
val less_pats = prep_pats ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   185
val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   186
4309
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   187
val nat_cancel_sums = map prep_simproc
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   188
  [("nateq_cancel_sums", eq_pats, EqCancelSums.proc),
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   189
   ("natless_cancel_sums", less_pats, LessCancelSums.proc),
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   190
   ("natle_cancel_sums", le_pats, LeCancelSums.proc)];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   191
4309
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   192
val nat_cancel_factor = map prep_simproc
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   193
  [("nateq_cancel_factor", eq_pats, EqCancelFactor.proc),
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   194
   ("natless_cancel_factor", less_pats, LessCancelFactor.proc),
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   195
   ("natle_cancel_factor", le_pats, LeCancelFactor.proc)];
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   196
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   197
val nat_cancel = nat_cancel_factor @ nat_cancel_sums;
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   198
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   199
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   200
end;