src/HOL/arith_data.ML
author wenzelm
Tue, 11 Jul 2006 12:16:52 +0200
changeset 20070 3f31fb81b83a
parent 20044 92cc2f4c7335
child 20217 25b068a99d2b
permissions -rw-r--r--
let_simproc: activate Variable.import;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/arith_data.ML
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, Stefan Berghofer and Tobias Nipkow
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     4
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     5
Various arithmetic proof procedures.
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     6
*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     7
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     8
(*---------------------------------------------------------------------------*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     9
(* 1. Cancellation of common terms                                           *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    10
(*---------------------------------------------------------------------------*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    11
13517
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    12
structure NatArithUtils =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    13
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    14
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    15
(** abstract syntax of structure nat: 0, Suc, + **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    16
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    17
(* mk_sum, mk_norm_sum *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    18
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    19
val one = HOLogic.mk_nat 1;
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
    20
val mk_plus = HOLogic.mk_binop "HOL.plus";
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    21
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    22
fun mk_sum [] = HOLogic.zero
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    23
  | mk_sum [t] = t
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    24
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    25
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    26
(*normal form of sums: Suc (... (Suc (a + (b + ...))))*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    27
fun mk_norm_sum ts =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    28
  let val (ones, sums) = List.partition (equal one) ts in
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    29
    funpow (length ones) HOLogic.mk_Suc (mk_sum sums)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    30
  end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    31
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    32
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    33
(* dest_sum *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    34
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
    35
val dest_plus = HOLogic.dest_bin "HOL.plus" HOLogic.natT;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    36
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    37
fun dest_sum tm =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    38
  if HOLogic.is_zero tm then []
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    39
  else
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    40
    (case try HOLogic.dest_Suc tm of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
    41
      SOME t => one :: dest_sum t
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
    42
    | NONE =>
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    43
        (case try dest_plus tm of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
    44
          SOME (t, u) => dest_sum t @ dest_sum u
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
    45
        | NONE => [tm]));
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    46
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    47
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    48
(** generic proof tools **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    49
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    50
(* prove conversions *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    51
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
    52
fun prove_conv expand_tac norm_tac ss tu =  (* FIXME avoid standard *)
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
    53
  mk_meta_eq (standard (Goal.prove (Simplifier.the_context ss) [] []
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
    54
      (HOLogic.mk_Trueprop (HOLogic.mk_eq tu))
17989
wenzelm
parents: 17985
diff changeset
    55
    (K (EVERY [expand_tac, norm_tac ss]))));
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    56
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    57
val subst_equals = prove_goal HOL.thy "[| t = s; u = t |] ==> u = s"
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    58
  (fn prems => [cut_facts_tac prems 1, SIMPSET' asm_simp_tac 1]);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    59
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    60
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    61
(* rewriting *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    62
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 17989
diff changeset
    63
fun simp_all_tac rules =
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 17989
diff changeset
    64
  let val ss0 = HOL_ss addsimps rules
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 17989
diff changeset
    65
  in fn ss => ALLGOALS (simp_tac (Simplifier.inherit_context ss ss0)) end;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    66
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    67
val add_rules = [add_Suc, add_Suc_right, add_0, add_0_right];
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    68
val mult_rules = [mult_Suc, mult_Suc_right, mult_0, mult_0_right];
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    69
13517
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    70
fun prep_simproc (name, pats, proc) =
16834
71d87aeebb57 HOL.Not;
wenzelm
parents: 16733
diff changeset
    71
  Simplifier.simproc (the_context ()) name pats proc;
13517
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    72
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    73
end;
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    74
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    75
signature ARITH_DATA =
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    76
sig
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    77
  val nat_cancel_sums_add: simproc list
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    78
  val nat_cancel_sums: simproc list
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    79
end;
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    80
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    81
structure ArithData: ARITH_DATA =
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    82
struct
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    83
42efec18f5b2 Added div+mod cancelling simproc
nipkow
parents: 13499
diff changeset
    84
open NatArithUtils;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    85
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    86
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    87
(** cancel common summands **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    88
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    89
structure Sum =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    90
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    91
  val mk_sum = mk_norm_sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    92
  val dest_sum = dest_sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    93
  val prove_conv = prove_conv;
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 17989
diff changeset
    94
  val norm_tac1 = simp_all_tac add_rules;
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 17989
diff changeset
    95
  val norm_tac2 = simp_all_tac add_ac;
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 17989
diff changeset
    96
  fun norm_tac ss = norm_tac1 ss THEN norm_tac2 ss;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    97
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    98
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    99
fun gen_uncancel_tac rule ct =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   100
  rtac (instantiate' [] [NONE, SOME ct] (rule RS subst_equals)) 1;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   101
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   102
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   103
(* nat eq *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   104
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   105
structure EqCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   106
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   107
  open Sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   108
  val mk_bal = HOLogic.mk_eq;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   109
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 13902
diff changeset
   110
  val uncancel_tac = gen_uncancel_tac nat_add_left_cancel;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   111
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   112
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   113
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   114
(* nat less *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   115
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   116
structure LessCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   117
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   118
  open Sum;
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   119
  val mk_bal = HOLogic.mk_binrel "Orderings.less";
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   120
  val dest_bal = HOLogic.dest_bin "Orderings.less" HOLogic.natT;
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 13902
diff changeset
   121
  val uncancel_tac = gen_uncancel_tac nat_add_left_cancel_less;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   122
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   123
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   124
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   125
(* nat le *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   126
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   127
structure LeCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   128
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   129
  open Sum;
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   130
  val mk_bal = HOLogic.mk_binrel "Orderings.less_eq";
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   131
  val dest_bal = HOLogic.dest_bin "Orderings.less_eq" HOLogic.natT;
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 13902
diff changeset
   132
  val uncancel_tac = gen_uncancel_tac nat_add_left_cancel_le;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   133
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   134
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   135
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   136
(* nat diff *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   137
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   138
structure DiffCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   139
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   140
  open Sum;
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   141
  val mk_bal = HOLogic.mk_binop "HOL.minus";
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   142
  val dest_bal = HOLogic.dest_bin "HOL.minus" HOLogic.natT;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   143
  val uncancel_tac = gen_uncancel_tac diff_cancel;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   144
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   145
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   146
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   147
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   148
(** prepare nat_cancel simprocs **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   149
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   150
val nat_cancel_sums_add = map prep_simproc
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12949
diff changeset
   151
  [("nateq_cancel_sums",
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
   152
     ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"], K EqCancelSums.proc),
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12949
diff changeset
   153
   ("natless_cancel_sums",
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
   154
     ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"], K LessCancelSums.proc),
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12949
diff changeset
   155
   ("natle_cancel_sums",
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
   156
     ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"], K LeCancelSums.proc)];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   157
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   158
val nat_cancel_sums = nat_cancel_sums_add @
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12949
diff changeset
   159
  [prep_simproc ("natdiff_cancel_sums",
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19823
diff changeset
   160
    ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"], K DiffCancelSums.proc)];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   161
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   162
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   163
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   164
open ArithData;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   165
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   166
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   167
(*---------------------------------------------------------------------------*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   168
(* 2. Linear arithmetic                                                      *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   169
(*---------------------------------------------------------------------------*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   170
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   171
(* Parameters data for general linear arithmetic functor *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   172
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   173
structure LA_Logic: LIN_ARITH_LOGIC =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   174
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   175
val ccontr = ccontr;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   176
val conjI = conjI;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   177
val notI = notI;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   178
val sym = sym;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   179
val not_lessD = linorder_not_less RS iffD1;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   180
val not_leD = linorder_not_le RS iffD1;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   181
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   182
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   183
fun mk_Eq thm = (thm RS Eq_FalseI) handle THM _ => (thm RS Eq_TrueI);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   184
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   185
val mk_Trueprop = HOLogic.mk_Trueprop;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   186
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16485
diff changeset
   187
fun atomize thm = case #prop(rep_thm thm) of
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16485
diff changeset
   188
    Const("Trueprop",_) $ (Const("op &",_) $ _ $ _) =>
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16485
diff changeset
   189
    atomize(thm RS conjunct1) @ atomize(thm RS conjunct2)
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16485
diff changeset
   190
  | _ => [thm];
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16485
diff changeset
   191
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   192
fun neg_prop(TP$(Const("Not",_)$t)) = TP$t
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   193
  | neg_prop(TP$t) = TP $ (Const("Not",HOLogic.boolT-->HOLogic.boolT)$t);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   194
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   195
fun is_False thm =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   196
  let val _ $ t = #prop(rep_thm thm)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   197
  in t = Const("False",HOLogic.boolT) end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   198
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   199
fun is_nat(t) = fastype_of1 t = HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   200
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   201
fun mk_nat_thm sg t =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   202
  let val ct = cterm_of sg t  and cn = cterm_of sg (Var(("n",0),HOLogic.natT))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   203
  in instantiate ([],[(cn,ct)]) le0 end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   204
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   205
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   206
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   207
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   208
(* arith theory data *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   209
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   210
structure ArithTheoryData = TheoryDataFun
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   211
(struct
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   212
  val name = "HOL/arith";
15185
8c43ffe2bb32 tuned "discrete" field
nipkow
parents: 15184
diff changeset
   213
  type T = {splits: thm list, inj_consts: (string * typ)list, discrete: string  list, presburger: (int -> tactic) option};
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   214
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   215
  val empty = {splits = [], inj_consts = [], discrete = [], presburger = NONE};
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   216
  val copy = I;
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   217
  val extend = I;
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   218
  fun merge _ ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, presburger= presburger1},
13877
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   219
             {splits= splits2, inj_consts= inj_consts2, discrete= discrete2, presburger= presburger2}) =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   220
   {splits = Drule.merge_rules (splits1, splits2),
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   221
    inj_consts = merge_lists inj_consts1 inj_consts2,
15185
8c43ffe2bb32 tuned "discrete" field
nipkow
parents: 15184
diff changeset
   222
    discrete = merge_lists discrete1 discrete2,
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   223
    presburger = (case presburger1 of NONE => presburger2 | p => p)};
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   224
  fun print _ _ = ();
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   225
end);
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   226
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   227
val arith_split_add = Thm.declaration_attribute (fn thm =>
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   228
  Context.map_theory (ArithTheoryData.map (fn {splits,inj_consts,discrete,presburger} =>
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   229
    {splits= thm::splits, inj_consts= inj_consts, discrete= discrete, presburger= presburger})));
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   230
13877
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   231
fun arith_discrete d = ArithTheoryData.map (fn {splits,inj_consts,discrete,presburger} =>
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   232
  {splits = splits, inj_consts = inj_consts, discrete = d :: discrete, presburger= presburger});
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   233
13877
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   234
fun arith_inj_const c = ArithTheoryData.map (fn {splits,inj_consts,discrete,presburger} =>
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   235
  {splits = splits, inj_consts = c :: inj_consts, discrete = discrete, presburger = presburger});
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   236
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   237
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   238
structure LA_Data_Ref: LIN_ARITH_DATA =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   239
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   240
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   241
(* Decomposition of terms *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   242
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   243
fun nT (Type("fun",[N,_])) = N = HOLogic.natT
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   244
  | nT _ = false;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   245
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 16970
diff changeset
   246
fun add_atom(t,m,(p,i)) = (case AList.lookup (op =) p t of NONE => ((t, m) :: p, i)
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   247
                           | SOME n => (AList.update (op =) (t, Rat.add (n, m)) p, i));
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   248
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   249
exception Zero;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   250
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   251
fun rat_of_term (numt,dent) =
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   252
  let val num = HOLogic.dest_binum numt and den = HOLogic.dest_binum dent
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   253
  in if den = 0 then raise Zero else Rat.rat_of_quotient (num,den) end;
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   254
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   255
(* Warning: in rare cases number_of encloses a non-numeral,
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   256
   in which case dest_binum raises TERM; hence all the handles below.
11334
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   257
   Same for Suc-terms that turn out not to be numerals -
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   258
   although the simplifier should eliminate those anyway...
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   259
*)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   260
11334
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   261
fun number_of_Sucs (Const("Suc",_) $ n) = number_of_Sucs n + 1
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   262
  | number_of_Sucs t = if HOLogic.is_zero t then 0
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   263
                       else raise TERM("number_of_Sucs",[])
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   264
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   265
(* decompose nested multiplications, bracketing them to the right and combining all
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   266
   their coefficients
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   267
*)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   268
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   269
fun demult inj_consts =
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   270
let
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   271
fun demult((mC as Const("HOL.times",_)) $ s $ t,m) = ((case s of
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   272
        Const("Numeral.number_of",_)$n
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   273
        => demult(t,Rat.mult(m,Rat.rat_of_intinf(HOLogic.dest_binum n)))
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   274
      | Const("HOL.uminus",_)$(Const("Numeral.number_of",_)$n)
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   275
        => demult(t,Rat.mult(m,Rat.rat_of_intinf(~(HOLogic.dest_binum n))))
11334
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   276
      | Const("Suc",_) $ _
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   277
        => demult(t,Rat.mult(m,Rat.rat_of_int(number_of_Sucs s)))
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   278
      | Const("HOL.times",_) $ s1 $ s2 => demult(mC $ s1 $ (mC $ s2 $ t),m)
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   279
      | Const("HOL.divide",_) $ numt $ (Const("Numeral.number_of",_)$dent) =>
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   280
          let val den = HOLogic.dest_binum dent
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   281
          in if den = 0 then raise Zero
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   282
             else demult(mC $ numt $ t,Rat.mult(m, Rat.inv(Rat.rat_of_intinf den)))
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   283
          end
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   284
      | _ => atomult(mC,s,t,m)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   285
      ) handle TERM _ => atomult(mC,s,t,m))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   286
  | demult(atom as Const("HOL.divide",_) $ t $ (Const("Numeral.number_of",_)$dent), m) =
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   287
      (let val den = HOLogic.dest_binum dent
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   288
       in if den = 0 then raise Zero else demult(t,Rat.mult(m, Rat.inv(Rat.rat_of_intinf den))) end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   289
       handle TERM _ => (SOME atom,m))
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   290
  | demult(Const("0",_),m) = (NONE, Rat.rat_of_int 0)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   291
  | demult(Const("1",_),m) = (NONE, m)
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   292
  | demult(t as Const("Numeral.number_of",_)$n,m) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   293
      ((NONE,Rat.mult(m,Rat.rat_of_intinf(HOLogic.dest_binum n)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   294
       handle TERM _ => (SOME t,m))
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   295
  | demult(Const("HOL.uminus",_)$t, m) = demult(t,Rat.mult(m,Rat.rat_of_int(~1)))
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   296
  | demult(t as Const f $ x, m) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   297
      (if f mem inj_consts then SOME x else SOME t,m)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   298
  | demult(atom,m) = (SOME atom,m)
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   299
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   300
and atomult(mC,atom,t,m) = (case demult(t,m) of (NONE,m') => (SOME atom,m')
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   301
                            | (SOME t',m') => (SOME(mC $ atom $ t'),m'))
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   302
in demult end;
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   303
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   304
fun decomp2 inj_consts (rel,lhs,rhs) =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   305
let
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   306
(* Turn term into list of summand * multiplicity plus a constant *)
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   307
fun poly(Const("HOL.plus",_) $ s $ t, m, pi) = poly(s,m,poly(t,m,pi))
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   308
  | poly(all as Const("HOL.minus",T) $ s $ t, m, pi) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   309
      if nT T then add_atom(all,m,pi) else poly(s,m,poly(t,Rat.neg m,pi))
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   310
  | poly(all as Const("HOL.uminus",T) $ t, m, pi) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   311
      if nT T then add_atom(all,m,pi) else poly(t,Rat.neg m,pi)
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   312
  | poly(Const("0",_), _, pi) = pi
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   313
  | poly(Const("1",_), m, (p,i)) = (p,Rat.add(i,m))
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   314
  | poly(Const("Suc",_)$t, m, (p,i)) = poly(t, m, (p,Rat.add(i,m)))
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 19043
diff changeset
   315
  | poly(t as Const("HOL.times",_) $ _ $ _, m, pi as (p,i)) =
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   316
      (case demult inj_consts (t,m) of
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   317
         (NONE,m') => (p,Rat.add(i,m))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   318
       | (SOME u,m') => add_atom(u,m',pi))
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   319
  | poly(t as Const("HOL.divide",_) $ _ $ _, m, pi as (p,i)) =
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   320
      (case demult inj_consts (t,m) of
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   321
         (NONE,m') => (p,Rat.add(i,m'))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   322
       | (SOME u,m') => add_atom(u,m',pi))
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   323
  | poly(all as (Const("Numeral.number_of",_)$t,m,(p,i))) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   324
      ((p,Rat.add(i,Rat.mult(m,Rat.rat_of_intinf(HOLogic.dest_binum t))))
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   325
       handle TERM _ => add_atom all)
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   326
  | poly(all as Const f $ x, m, pi) =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   327
      if f mem inj_consts then poly(x,m,pi) else add_atom(all,m,pi)
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   328
  | poly x  = add_atom x;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   329
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   330
val (p,i) = poly(lhs,Rat.rat_of_int 1,([],Rat.rat_of_int 0))
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17875
diff changeset
   331
and (q,j) = poly(rhs,Rat.rat_of_int 1,([],Rat.rat_of_int 0))
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   332
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   333
  in case rel of
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   334
       "Orderings.less"  => SOME(p,i,"<",q,j)
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   335
     | "Orderings.less_eq" => SOME(p,i,"<=",q,j)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   336
     | "op ="  => SOME(p,i,"=",q,j)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   337
     | _       => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   338
  end handle Zero => NONE;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   339
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   340
fun negate(SOME(x,i,rel,y,j,d)) = SOME(x,i,"~"^rel,y,j,d)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   341
  | negate NONE = NONE;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   342
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   343
fun of_lin_arith_sort sg U =
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   344
  Type.of_sort (Sign.tsig_of sg) (U,["Ring_and_Field.ordered_idom"])
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   345
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   346
fun allows_lin_arith sg discrete (U as Type(D,[])) =
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   347
      if of_lin_arith_sort sg U
15185
8c43ffe2bb32 tuned "discrete" field
nipkow
parents: 15184
diff changeset
   348
      then (true, D mem discrete)
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   349
      else (* special cases *)
15185
8c43ffe2bb32 tuned "discrete" field
nipkow
parents: 15184
diff changeset
   350
           if D mem discrete then (true,true) else (false,false)
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   351
  | allows_lin_arith sg discrete U = (of_lin_arith_sort sg U, false);
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   352
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   353
fun decomp1 (sg,discrete,inj_consts) (T,xxx) =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   354
  (case T of
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   355
     Type("fun",[U,_]) =>
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   356
       (case allows_lin_arith sg discrete U of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   357
          (true,d) => (case decomp2 inj_consts xxx of NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   358
                       | SOME(p,i,rel,q,j) => SOME(p,i,rel,q,j,d))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   359
        | (false,_) => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   360
   | _ => NONE);
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   361
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   362
fun decomp2 data (_$(Const(rel,T)$lhs$rhs)) = decomp1 data (T,(rel,lhs,rhs))
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   363
  | decomp2 data (_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   364
      negate(decomp1 data (T,(rel,lhs,rhs)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   365
  | decomp2 data _ = NONE
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   366
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   367
fun decomp sg =
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   368
  let val {discrete, inj_consts, ...} = ArithTheoryData.get sg
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   369
  in decomp2 (sg,discrete,inj_consts) end
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   370
19481
a6205c6203ea renamed HOLogic.mk_bin to mk_binum for consistency with dest_binum
paulson
parents: 19297
diff changeset
   371
fun number_of(n,T) = HOLogic.number_of_const T $ (HOLogic.mk_binum n)
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   372
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   373
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   374
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   375
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   376
structure Fast_Arith =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   377
  Fast_Lin_Arith(structure LA_Logic=LA_Logic and LA_Data=LA_Data_Ref);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   378
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   379
val fast_arith_tac    = Fast_Arith.lin_arith_tac false
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   380
and fast_ex_arith_tac = Fast_Arith.lin_arith_tac
14517
7ae3b247c6e9 exposed fast_arith_neq_limit
nipkow
parents: 14509
diff changeset
   381
and trace_arith    = Fast_Arith.trace
7ae3b247c6e9 exposed fast_arith_neq_limit
nipkow
parents: 14509
diff changeset
   382
and fast_arith_neq_limit = Fast_Arith.fast_arith_neq_limit;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   383
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   384
local
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   385
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   386
(* reduce contradictory <= to False.
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   387
   Most of the work is done by the cancel tactics.
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   388
*)
12931
2c0251fada94 solved the problem that Larry's simproce cancle_numerals(?) returns
nipkow
parents: 12480
diff changeset
   389
val add_rules =
14368
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14356
diff changeset
   390
 [add_zero_left,add_zero_right,Zero_not_Suc,Suc_not_Zero,le_0_eq,
19297
8f6e097d7b23 Removal of unnecessary simprules: simproc cancel_numerals now works without
paulson
parents: 19285
diff changeset
   391
  One_nat_def,
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17611
diff changeset
   392
  order_less_irrefl, zero_neq_one, zero_less_one, zero_le_one,
16473
b24c820a0b85 moving some generic inequalities from integer arith to nat arith
paulson
parents: 16424
diff changeset
   393
  zero_neq_one RS not_sym, not_one_le_zero, not_one_less_zero];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   394
14368
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14356
diff changeset
   395
val add_mono_thms_ordered_semiring = map (fn s => prove_goal (the_context ()) s
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   396
 (fn prems => [cut_facts_tac prems 1,
14368
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14356
diff changeset
   397
               blast_tac (claset() addIs [add_mono]) 1]))
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   398
["(i <= j) & (k <= l) ==> i + k <= j + (l::'a::pordered_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   399
 "(i  = j) & (k <= l) ==> i + k <= j + (l::'a::pordered_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   400
 "(i <= j) & (k  = l) ==> i + k <= j + (l::'a::pordered_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   401
 "(i  = j) & (k  = l) ==> i + k  = j + (l::'a::pordered_ab_semigroup_add)"
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   402
];
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   403
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   404
val mono_ss = simpset() addsimps
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   405
                [add_mono,add_strict_mono,add_less_le_mono,add_le_less_mono];
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   406
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   407
val add_mono_thms_ordered_field =
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   408
  map (fn s => prove_goal (the_context ()) s
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   409
                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   410
    ["(i<j) & (k=l)   ==> i+k < j+(l::'a::pordered_cancel_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   411
     "(i=j) & (k<l)   ==> i+k < j+(l::'a::pordered_cancel_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   412
     "(i<j) & (k<=l)  ==> i+k < j+(l::'a::pordered_cancel_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   413
     "(i<=j) & (k<l)  ==> i+k < j+(l::'a::pordered_cancel_ab_semigroup_add)",
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   414
     "(i<j) & (k<l)   ==> i+k < j+(l::'a::pordered_cancel_ab_semigroup_add)"];
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   415
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   416
in
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   417
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   418
val init_lin_arith_data =
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   419
 Fast_Arith.setup #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   420
 Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, ...} =>
15121
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   421
   {add_mono_thms = add_mono_thms @
1198032bad25 Initial changes to extend arithmetic from individual types to type classes.
nipkow
parents: 14738
diff changeset
   422
    add_mono_thms_ordered_semiring @ add_mono_thms_ordered_field,
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   423
    mult_mono_thms = mult_mono_thms,
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   424
    inj_thms = inj_thms,
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   425
    lessD = lessD @ [Suc_leI],
15923
01d5d0c1c078 fixed lin.arith
nipkow
parents: 15921
diff changeset
   426
    neqE = [linorder_neqE_nat,
16485
77ae3bfa8b76 get_thm instead of obsolete Goals.get_thm;
wenzelm
parents: 16473
diff changeset
   427
      get_thm (theory "Ring_and_Field") (Name "linorder_neqE_ordered_idom")],
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15221
diff changeset
   428
    simpset = HOL_basic_ss addsimps add_rules
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17611
diff changeset
   429
                   addsimprocs [ab_group_add_cancel.sum_conv,
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15221
diff changeset
   430
                                ab_group_add_cancel.rel_conv]
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15221
diff changeset
   431
                   (*abel_cancel helps it work in abstract algebraic domains*)
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   432
                   addsimprocs nat_cancel_sums_add}) #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   433
  ArithTheoryData.init #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   434
  arith_discrete "nat";
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   435
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   436
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   437
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12949
diff changeset
   438
val fast_nat_arith_simproc =
16834
71d87aeebb57 HOL.Not;
wenzelm
parents: 16733
diff changeset
   439
  Simplifier.simproc (the_context ()) "fast_nat_arith"
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12949
diff changeset
   440
    ["(m::nat) < n","(m::nat) <= n", "(m::nat) = n"] Fast_Arith.lin_arith_prover;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   441
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   442
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   443
(* Because of fast_nat_arith_simproc, the arithmetic solver is really only
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   444
useful to detect inconsistencies among the premises for subgoals which are
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   445
*not* themselves (in)equalities, because the latter activate
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   446
fast_nat_arith_simproc anyway. However, it seems cheaper to activate the
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   447
solver all the time rather than add the additional check. *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   448
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   449
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   450
(* arith proof method *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   451
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   452
(* FIXME: K true should be replaced by a sensible test to speed things up
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   453
   in case there are lots of irrelevant terms involved;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   454
   elimination of min/max can be optimized:
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   455
   (max m n + k <= r) = (m+k <= r & n+k <= r)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   456
   (l <= min m n + k) = (l <= m+k & l <= n+k)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   457
*)
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   458
local
19043
6c0fca729f33 typo in a comment fixed
webertj
parents: 18728
diff changeset
   459
(* a simpset for computations subject to optimization !!! *)
18394
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   460
(*
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   461
val binarith = map thm
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   462
  ["Pls_0_eq", "Min_1_eq",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   463
 "bin_pred_Pls","bin_pred_Min","bin_pred_1","bin_pred_0",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   464
  "bin_succ_Pls", "bin_succ_Min", "bin_succ_1", "bin_succ_0",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   465
  "bin_add_Pls", "bin_add_Min", "bin_add_BIT_0", "bin_add_BIT_10",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   466
  "bin_add_BIT_11", "bin_minus_Pls", "bin_minus_Min", "bin_minus_1", 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   467
  "bin_minus_0", "bin_mult_Pls", "bin_mult_Min", "bin_mult_1", "bin_mult_0", 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   468
  "bin_add_Pls_right", "bin_add_Min_right"];
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   469
 val intarithrel = 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   470
     (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   471
		"int_le_number_of_eq","int_iszero_number_of_0",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   472
		"int_less_number_of_eq_neg"]) @
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   473
     (map (fn s => thm s RS thm "lift_bool") 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   474
	  ["int_iszero_number_of_Pls","int_iszero_number_of_1",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   475
	   "int_neg_number_of_Min"])@
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   476
     (map (fn s => thm s RS thm "nlift_bool") 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   477
	  ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]);
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   478
     
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   479
val intarith = map thm ["int_number_of_add_sym", "int_number_of_minus_sym",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   480
			"int_number_of_diff_sym", "int_number_of_mult_sym"];
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   481
val natarith = map thm ["add_nat_number_of", "diff_nat_number_of",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   482
			"mult_nat_number_of", "eq_nat_number_of",
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   483
			"less_nat_number_of"]
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   484
val powerarith = 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   485
    (map thm ["nat_number_of", "zpower_number_of_even", 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   486
	      "zpower_Pls", "zpower_Min"]) @ 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   487
    [(Tactic.simplify true [thm "zero_eq_Numeral0_nring", 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   488
			   thm "one_eq_Numeral1_nring"] 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   489
  (thm "zpower_number_of_odd"))]
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   490
18394
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   491
val comp_arith = binarith @ intarith @ intarithrel @ natarith 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   492
	    @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"];
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   493
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   494
val comp_ss = HOL_basic_ss addsimps comp_arith addsimps simp_thms;
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   495
*)
13499
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   496
fun raw_arith_tac ex i st =
f95f5818f24f Counter example generation mods.
nipkow
parents: 13462
diff changeset
   497
  refute_tac (K true)
19823
9e4573eaacb3 replaced REPEAT by REPOEAT_DETERM
nipkow
parents: 19481
diff changeset
   498
   (REPEAT_DETERM o split_tac (#splits (ArithTheoryData.get (Thm.theory_of_thm st))))
18394
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   499
(*   (REPEAT o 
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   500
    (fn i =>(split_tac (#splits (ArithTheoryData.get(Thm.theory_of_thm st))) i)
fa0674cd6df1 simpset for computation in raw_arith_tac added just as comment, nothing changed!
chaieb
parents: 18328
diff changeset
   501
		THEN (simp_tac comp_ss i))) *)
14509
9d8978a2ae56 got rid of ignore_neq again.
nipkow
parents: 14507
diff changeset
   502
   ((REPEAT_DETERM o etac linorder_neqE) THEN' fast_ex_arith_tac ex)
9d8978a2ae56 got rid of ignore_neq again.
nipkow
parents: 14507
diff changeset
   503
   i st;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   504
13877
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   505
fun presburger_tac i st =
16834
71d87aeebb57 HOL.Not;
wenzelm
parents: 16733
diff changeset
   506
  (case ArithTheoryData.get (Thm.theory_of_thm st) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   507
     {presburger = SOME tac, ...} =>
16970
c1ef99e08c39 tuned msg;
wenzelm
parents: 16834
diff changeset
   508
       (warning "Trying full Presburger arithmetic ..."; tac i st)
13877
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   509
   | _ => no_tac st);
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   510
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   511
in
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   512
13877
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   513
val simple_arith_tac = FIRST' [fast_arith_tac,
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   514
  ObjectLogic.atomize_tac THEN' raw_arith_tac true];
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   515
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   516
val arith_tac = FIRST' [fast_arith_tac,
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   517
  ObjectLogic.atomize_tac THEN' raw_arith_tac true,
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   518
  presburger_tac];
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   519
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   520
val silent_arith_tac = FIRST' [fast_arith_tac,
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   521
  ObjectLogic.atomize_tac THEN' raw_arith_tac false,
a6b825ee48d9 Added hook for presburger arithmetic decision procedure.
berghofe
parents: 13517
diff changeset
   522
  presburger_tac];
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   523
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   524
fun arith_method prems =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   525
  Method.METHOD (fn facts => HEADGOAL (Method.insert_tac (prems @ facts) THEN' arith_tac));
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   526
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   527
end;
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   528
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   529
(* antisymmetry:
15197
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15195
diff changeset
   530
   combines x <= y (or ~(y < x)) and y <= x (or ~(x < y)) into x = y
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   531
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   532
local
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   533
val antisym = mk_meta_eq order_antisym
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   534
val not_lessD = linorder_not_less RS iffD1
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   535
fun prp t thm = (#prop(rep_thm thm) = t)
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   536
in
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   537
fun antisym_eq prems thm =
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   538
  let
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   539
    val r = #prop(rep_thm thm);
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   540
  in
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   541
    case r of
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   542
      Tr $ ((c as Const("Orderings.less_eq",T)) $ s $ t) =>
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   543
        let val r' = Tr $ (c $ t $ s)
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   544
        in
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   545
          case Library.find_first (prp r') prems of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   546
            NONE =>
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   547
              let val r' = Tr $ (HOLogic.Not $ (Const("Orderings.less",T) $ s $ t))
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   548
              in case Library.find_first (prp r') prems of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   549
                   NONE => []
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   550
                 | SOME thm' => [(thm' RS not_lessD) RS (thm RS antisym)]
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   551
              end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   552
          | SOME thm' => [thm' RS (thm RS antisym)]
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   553
        end
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   554
    | Tr $ (Const("Not",_) $ (Const("Orderings.less",T) $ s $ t)) =>
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   555
        let val r' = Tr $ (Const("Orderings.less_eq",T) $ s $ t)
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   556
        in
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   557
          case Library.find_first (prp r') prems of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   558
            NONE =>
19277
f7602e74d948 renamed op < <= to Orderings.less(_eq)
haftmann
parents: 19233
diff changeset
   559
              let val r' = Tr $ (HOLogic.Not $ (Const("Orderings.less",T) $ t $ s))
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   560
              in case Library.find_first (prp r') prems of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   561
                   NONE => []
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   562
                 | SOME thm' =>
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   563
                     [(thm' RS not_lessD) RS ((thm RS not_lessD) RS antisym)]
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   564
              end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15234
diff changeset
   565
          | SOME thm' => [thm' RS ((thm RS not_lessD) RS antisym)]
15195
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   566
        end
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   567
    | _ => []
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   568
  end
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   569
  handle THM _ => []
197e00ce3f20 new forward deduction capability for simplifier
nipkow
parents: 15185
diff changeset
   570
end;
15197
19e735596e51 Added antisymmetry simproc
nipkow
parents: 15195
diff changeset
   571
*)
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   572
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   573
(* theory setup *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   574
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   575
val arith_setup =
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   576
  init_lin_arith_data #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   577
  (fn thy => (Simplifier.change_simpset_of thy (fn ss => ss
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17611
diff changeset
   578
    addsimprocs (nat_cancel_sums @ [fast_nat_arith_simproc])
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   579
    addSolver (mk_solver' "lin. arith." Fast_Arith.cut_lin_arith_tac)); thy)) #>
15221
8412cfdf3287 tweaking of arithmetic proofs
paulson
parents: 15197
diff changeset
   580
  Method.add_methods
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17611
diff changeset
   581
    [("arith", (arith_method o #2) oo Method.syntax Args.bang_facts,
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   582
      "decide linear arithmethic")] #>
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   583
  Attrib.add_attributes [("arith_split", Attrib.no_args arith_split_add,
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18394
diff changeset
   584
    "declaration of split rules for arithmetic procedure")];