src/HOL/arith_data.ML
author wenzelm
Sun, 12 Jul 1998 11:49:17 +0200
changeset 5132 24f992a25adc
parent 4675 6efc56450d09
child 5303 22029546d109
permissions -rw-r--r--
isatool expandshort;
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
4310
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    21
(* mk_sum, mk_norm_sum *)
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    22
4310
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    23
val one = HOLogic.mk_nat 1;
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    24
val mk_plus = HOLogic.mk_binop "op +";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    25
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    26
fun mk_sum [] = HOLogic.zero
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    27
  | mk_sum [t] = t
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    28
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    29
4310
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    30
(*normal form of sums: Suc (... (Suc (a + (b + ...))))*)
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    31
fun mk_norm_sum ts =
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    32
  let val (ones, sums) = partition (equal one) ts in
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    33
    funpow (length ones) HOLogic.mk_Suc (mk_sum sums)
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    34
  end;
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    35
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    36
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    37
(* dest_sum *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    38
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    39
val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    40
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    41
fun dest_sum tm =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    42
  if HOLogic.is_zero tm then []
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    43
  else
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    44
    (case try HOLogic.dest_Suc tm of
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    45
      Some t => one :: dest_sum t
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    46
    | None =>
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    47
        (case try dest_plus tm of
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    48
          Some (t, u) => dest_sum t @ dest_sum u
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    49
        | None => [tm]));
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    50
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    51
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    52
(** generic proof tools **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    53
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    54
(* prove conversions *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    55
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    56
val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    57
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    58
fun prove_conv expand_tac norm_tac sg (t, u) =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    59
  mk_meta_eq (prove_goalw_cterm [] (cterm_of sg (mk_eqv (t, u)))
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    60
    (K [expand_tac, norm_tac]))
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    61
  handle ERROR => error ("The error(s) above occurred while trying to prove " ^
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    62
    (string_of_cterm (cterm_of sg (mk_eqv (t, u)))));
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    63
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    64
val subst_equals = prove_goal HOL.thy "[| t = s; u = t |] ==> u = s"
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    65
  (fn prems => [cut_facts_tac prems 1, SIMPSET' asm_simp_tac 1]);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    66
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    67
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    68
(* rewriting *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    69
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    70
fun simp_all rules = ALLGOALS (simp_tac (HOL_ss addsimps rules));
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    71
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    72
val add_rules = [add_Suc, add_Suc_right, add_0, add_0_right];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    73
val mult_rules = [mult_Suc, mult_Suc_right, mult_0, mult_0_right];
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    74
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    75
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    76
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    77
(** cancel common summands **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    78
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    79
structure Sum =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    80
struct
4310
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
    81
  val mk_sum = mk_norm_sum;
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    82
  val dest_sum = dest_sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    83
  val prove_conv = prove_conv;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    84
  val norm_tac = simp_all add_rules THEN simp_all add_ac;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    85
end;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    86
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    87
fun gen_uncancel_tac rule ct =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    88
  rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    89
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    90
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    91
(* nat eq *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    92
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    93
structure EqCancelSums = CancelSumsFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    94
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    95
  open Sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    96
  val mk_bal = HOLogic.mk_eq;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    97
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    98
  val uncancel_tac = gen_uncancel_tac add_left_cancel;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
    99
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   100
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   101
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   102
(* nat less *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   103
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   104
structure LessCancelSums = CancelSumsFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   105
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   106
  open Sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   107
  val mk_bal = HOLogic.mk_binrel "op <";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   108
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   109
  val uncancel_tac = gen_uncancel_tac add_left_cancel_less;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   110
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   111
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   112
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   113
(* nat le *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   114
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   115
structure LeCancelSums = CancelSumsFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   116
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   117
  open Sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   118
  val mk_bal = HOLogic.mk_binrel "op <=";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   119
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   120
  val uncancel_tac = gen_uncancel_tac add_left_cancel_le;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   121
end);
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
(* nat diff *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   125
4332
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   126
structure DiffCancelSums = CancelSumsFun
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   127
(struct
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   128
  open Sum;
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   129
  val mk_bal = HOLogic.mk_binop "op -";
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   130
  val dest_bal = HOLogic.dest_bin "op -" HOLogic.natT;
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   131
  val uncancel_tac = gen_uncancel_tac diff_cancel;
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   132
end);
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   133
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   134
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   135
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   136
(** cancel common factor **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   137
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   138
structure Factor =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   139
struct
4310
cad4f24be60b mk_norm_sum;
wenzelm
parents: 4309
diff changeset
   140
  val mk_sum = mk_norm_sum;
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   141
  val dest_sum = dest_sum;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   142
  val prove_conv = prove_conv;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   143
  val norm_tac = simp_all (add_rules @ mult_rules) THEN simp_all add_ac;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   144
end;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   145
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   146
fun mk_cnat n = cterm_of (sign_of Nat.thy) (HOLogic.mk_nat n);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   147
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   148
fun gen_multiply_tac rule k =
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   149
  if k > 0 then
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   150
    rtac (instantiate' [] [None, Some (mk_cnat (k - 1))] (rule RS subst_equals)) 1
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   151
  else no_tac;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   152
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   153
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   154
(* nat eq *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   155
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   156
structure EqCancelFactor = CancelFactorFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   157
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   158
  open Factor;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   159
  val mk_bal = HOLogic.mk_eq;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   160
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   161
  val multiply_tac = gen_multiply_tac Suc_mult_cancel1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   162
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   163
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   164
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   165
(* nat less *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   166
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   167
structure LessCancelFactor = CancelFactorFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   168
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   169
  open Factor;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   170
  val mk_bal = HOLogic.mk_binrel "op <";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   171
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   172
  val multiply_tac = gen_multiply_tac Suc_mult_less_cancel1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   173
end);
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
(* nat le *)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   177
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   178
structure LeCancelFactor = CancelFactorFun
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   179
(struct
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   180
  open Factor;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   181
  val mk_bal = HOLogic.mk_binrel "op <=";
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   182
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   183
  val multiply_tac = gen_multiply_tac Suc_mult_le_cancel1;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   184
end);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   185
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   186
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   187
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   188
(** prepare nat_cancel simprocs **)
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   189
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   190
fun prep_pat s = Thm.read_cterm (sign_of Arith.thy) (s, HOLogic.termTVar);
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   191
val prep_pats = map prep_pat;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   192
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   193
fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   194
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   195
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
   196
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
   197
val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"];
4332
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   198
val diff_pats = prep_pats ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"];
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   199
4309
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   200
val nat_cancel_sums = map prep_simproc
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   201
  [("nateq_cancel_sums", eq_pats, EqCancelSums.proc),
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   202
   ("natless_cancel_sums", less_pats, LessCancelSums.proc),
4332
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   203
   ("natle_cancel_sums", le_pats, LeCancelSums.proc),
d4a15e32c024 Added DiffCancelSums.
berghofe
parents: 4310
diff changeset
   204
   ("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   205
4309
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   206
val nat_cancel_factor = map prep_simproc
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   207
  [("nateq_cancel_factor", eq_pats, EqCancelFactor.proc),
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   208
   ("natless_cancel_factor", less_pats, LessCancelFactor.proc),
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   209
   ("natle_cancel_factor", le_pats, LeCancelFactor.proc)];
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   210
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   211
val nat_cancel = nat_cancel_factor @ nat_cancel_sums;
c98f44948d86 separate lists of simprocs;
wenzelm
parents: 4295
diff changeset
   212
4295
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   213
81132a38996a Setup various arithmetic proof procedures.
wenzelm
parents:
diff changeset
   214
end;
4336
wenzelm
parents: 4332
diff changeset
   215
wenzelm
parents: 4332
diff changeset
   216
wenzelm
parents: 4332
diff changeset
   217
open ArithData;
4368
1f2dd130fe39 simplification procedures nat_cancel enabled by default;
wenzelm
parents: 4336
diff changeset
   218
1f2dd130fe39 simplification procedures nat_cancel enabled by default;
wenzelm
parents: 4336
diff changeset
   219
1f2dd130fe39 simplification procedures nat_cancel enabled by default;
wenzelm
parents: 4336
diff changeset
   220
context Arith.thy;
1f2dd130fe39 simplification procedures nat_cancel enabled by default;
wenzelm
parents: 4336
diff changeset
   221
Addsimprocs nat_cancel;
4675
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   222
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   223
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   224
(*This proof requires natdiff_cancel_sums*)
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   225
goal Arith.thy "!!n::nat. m<n --> m<l --> (l-n) < (l-m)";
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   226
by (induct_tac "l" 1);
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   227
by (Simp_tac 1);
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   228
by (Clarify_tac 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 4675
diff changeset
   229
by (etac less_SucE 1);
4675
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   230
by (asm_simp_tac (simpset() addsimps [diff_Suc_le_Suc_diff RS le_less_trans,
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   231
				      Suc_diff_n]) 1);
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   232
by (Clarify_tac 1);
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   233
by (asm_simp_tac (simpset() addsimps [less_imp_diff_is_0]) 1);
6efc56450d09 New theorem
paulson
parents: 4368
diff changeset
   234
qed_spec_mp "diff_less_mono2";