src/HOL/arith_data.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 11334 a16eaf2a1edd
child 11701 3d51fbf81c17
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
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
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    12
signature ARITH_DATA =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    13
sig
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    14
  val nat_cancel_sums_add: simproc list
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    15
  val nat_cancel_sums: simproc list
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    16
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    17
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    18
structure ArithData: ARITH_DATA =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    19
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    20
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
(** abstract syntax of structure nat: 0, Suc, + **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    23
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    24
(* mk_sum, mk_norm_sum *)
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
val one = HOLogic.mk_nat 1;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    27
val mk_plus = HOLogic.mk_binop "op +";
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    28
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    29
fun mk_sum [] = HOLogic.zero
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    30
  | mk_sum [t] = t
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    31
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
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
(*normal form of sums: Suc (... (Suc (a + (b + ...))))*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    34
fun mk_norm_sum ts =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    35
  let val (ones, sums) = partition (equal one) ts in
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    36
    funpow (length ones) HOLogic.mk_Suc (mk_sum sums)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    37
  end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    38
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    39
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    40
(* dest_sum *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    41
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    42
val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    43
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    44
fun dest_sum tm =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    45
  if HOLogic.is_zero tm then []
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    46
  else
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    47
    (case try HOLogic.dest_Suc tm of
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    48
      Some t => one :: dest_sum t
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    49
    | None =>
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    50
        (case try dest_plus tm of
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    51
          Some (t, u) => dest_sum t @ dest_sum u
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    52
        | None => [tm]));
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    53
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    54
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    55
(** generic proof tools **)
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
(* prove conversions *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    58
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    59
val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq;
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
fun prove_conv expand_tac norm_tac sg (t, u) =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    62
  mk_meta_eq (prove_goalw_cterm_nocheck [] (cterm_of sg (mk_eqv (t, u)))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    63
    (K [expand_tac, norm_tac]))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    64
  handle ERROR => error ("The error(s) above occurred while trying to prove " ^
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    65
    (string_of_cterm (cterm_of sg (mk_eqv (t, u)))));
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 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
    68
  (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
    69
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    70
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    71
(* rewriting *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    72
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    73
fun simp_all rules = ALLGOALS (simp_tac (HOL_ss addsimps rules));
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    74
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    75
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
    76
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
    77
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    78
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    79
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    80
(** cancel common summands **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    81
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    82
structure Sum =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    83
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    84
  val mk_sum = mk_norm_sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    85
  val dest_sum = dest_sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    86
  val prove_conv = prove_conv;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    87
  val norm_tac = simp_all add_rules THEN simp_all add_ac;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    88
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    89
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    90
fun gen_uncancel_tac rule ct =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    91
  rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    92
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    93
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    94
(* nat eq *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    95
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    96
structure EqCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    97
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    98
  open Sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    99
  val mk_bal = HOLogic.mk_eq;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   100
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   101
  val uncancel_tac = gen_uncancel_tac add_left_cancel;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   102
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   103
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
(* nat less *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   106
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   107
structure LessCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   108
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   109
  open Sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   110
  val mk_bal = HOLogic.mk_binrel "op <";
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   111
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   112
  val uncancel_tac = gen_uncancel_tac add_left_cancel_less;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   113
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   114
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
(* nat le *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   117
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   118
structure LeCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   119
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   120
  open Sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   121
  val mk_bal = HOLogic.mk_binrel "op <=";
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   122
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   123
  val uncancel_tac = gen_uncancel_tac add_left_cancel_le;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   124
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   125
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
(* nat diff *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   128
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   129
structure DiffCancelSums = CancelSumsFun
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   130
(struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   131
  open Sum;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   132
  val mk_bal = HOLogic.mk_binop "op -";
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   133
  val dest_bal = HOLogic.dest_bin "op -" HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   134
  val uncancel_tac = gen_uncancel_tac diff_cancel;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   135
end);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   136
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
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   139
(** prepare nat_cancel simprocs **)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   140
10766
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   141
fun prep_pat s = Thm.read_cterm (Theory.sign_of (the_context ())) 
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   142
                                (s, HOLogic.termT);
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   143
val prep_pats = map prep_pat;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   144
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   145
fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   146
10766
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   147
val eq_pats = prep_pats ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", 
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   148
                         "m = Suc n"];
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   149
val less_pats = prep_pats ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n",
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   150
                           "m < Suc n"];
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   151
val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n", 
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   152
                         "Suc m <= n", "m <= Suc n"];
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   153
val diff_pats = prep_pats ["((l::nat) + m) - n", "(l::nat) - (m + n)", 
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   154
                           "Suc m - n", "m - Suc n"];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   155
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   156
val nat_cancel_sums_add = map prep_simproc
10766
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   157
  [("nateq_cancel_sums",   eq_pats,   EqCancelSums.proc),
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   158
   ("natless_cancel_sums", less_pats, LessCancelSums.proc),
10766
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   159
   ("natle_cancel_sums",   le_pats,   LeCancelSums.proc)];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   160
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   161
val nat_cancel_sums = nat_cancel_sums_add @
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   162
  [prep_simproc("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
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
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   165
end;
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
open ArithData;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   168
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
(* 2. Linear arithmetic                                                      *)
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
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   174
(* Parameters data for general linear arithmetic functor *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   175
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   176
structure LA_Logic: LIN_ARITH_LOGIC =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   177
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   178
val ccontr = ccontr;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   179
val conjI = conjI;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   180
val neqE = linorder_neqE;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   181
val notI = notI;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   182
val sym = sym;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   183
val not_lessD = linorder_not_less RS iffD1;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   184
val not_leD = linorder_not_le RS iffD1;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   185
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   186
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   187
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
   188
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   189
val mk_Trueprop = HOLogic.mk_Trueprop;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   190
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   191
fun neg_prop(TP$(Const("Not",_)$t)) = TP$t
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   192
  | 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
   193
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   194
fun is_False thm =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   195
  let val _ $ t = #prop(rep_thm thm)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   196
  in t = Const("False",HOLogic.boolT) end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   197
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   198
fun is_nat(t) = fastype_of1 t = HOLogic.natT;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   199
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   200
fun mk_nat_thm sg t =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   201
  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
   202
  in instantiate ([],[(cn,ct)]) le0 end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   203
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   204
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   205
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
(* arith theory data *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   208
9593
b732997cfc11 tuned names;
wenzelm
parents: 9436
diff changeset
   209
structure ArithTheoryDataArgs =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   210
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   211
  val name = "HOL/arith";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   212
  type T = {splits: thm list, inj_consts: (string * typ)list, discrete: (string * bool) list};
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   213
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   214
  val empty = {splits = [], inj_consts = [], discrete = []};
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   215
  val copy = I;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   216
  val prep_ext = I;
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   217
  fun merge ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1},
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   218
             {splits= splits2, inj_consts= inj_consts2, discrete= discrete2}) =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   219
   {splits = Drule.merge_rules (splits1, splits2),
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   220
    inj_consts = merge_lists inj_consts1 inj_consts2,
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   221
    discrete = merge_alists discrete1 discrete2};
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   222
  fun print _ _ = ();
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   223
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   224
9593
b732997cfc11 tuned names;
wenzelm
parents: 9436
diff changeset
   225
structure ArithTheoryData = TheoryDataFun(ArithTheoryDataArgs);
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   226
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   227
fun arith_split_add (thy, thm) = (ArithTheoryData.map (fn {splits,inj_consts,discrete} =>
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   228
  {splits= thm::splits, inj_consts= inj_consts, discrete= discrete}) thy, thm);
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   229
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   230
fun arith_discrete d = ArithTheoryData.map (fn {splits,inj_consts,discrete} =>
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   231
  {splits = splits, inj_consts = inj_consts, discrete = d :: discrete});
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   232
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   233
fun arith_inj_const c = ArithTheoryData.map (fn {splits,inj_consts,discrete} =>
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   234
  {splits = splits, inj_consts = c :: inj_consts, discrete = discrete});
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   235
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
structure LA_Data_Ref: LIN_ARITH_DATA =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   238
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   239
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   240
(* Decomposition of terms *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   241
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   242
fun nT (Type("fun",[N,_])) = N = HOLogic.natT
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   243
  | nT _ = false;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   244
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   245
fun add_atom(t,m,(p,i)) = (case assoc(p,t) of None => ((t,m)::p,i)
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   246
                           | Some n => (overwrite(p,(t,ratadd(n,m))), i));
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   247
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   248
exception Zero;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   249
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   250
fun rat_of_term(numt,dent) =
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   251
  let val num = HOLogic.dest_binum numt and den = HOLogic.dest_binum dent
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   252
  in if den = 0 then raise Zero else int_ratdiv(num,den) end;
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   253
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   254
(* Warning: in rare cases number_of encloses a non-numeral,
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   255
   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
   256
   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
   257
   although the simplifier should eliminate those anyway...
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   258
*)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   259
11334
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   260
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
   261
  | 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
   262
                       else raise TERM("number_of_Sucs",[])
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   263
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   264
(* decompose nested multiplications, bracketing them to the right and combining all
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   265
   their coefficients
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   266
*)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   267
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   268
fun demult((mC as Const("op *",_)) $ s $ t,m) = ((case s of
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   269
        Const("Numeral.number_of",_)$n
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   270
        => demult(t,ratmul(m,rat_of_int(HOLogic.dest_binum n)))
11334
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   271
      | Const("Suc",_) $ _
a16eaf2a1edd Allow Suc-numerals as coefficients in lin-arith formulae
nipkow
parents: 10906
diff changeset
   272
        => demult(t,ratmul(m,rat_of_int(number_of_Sucs s)))
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   273
      | Const("op *",_) $ s1 $ s2 => demult(mC $ s1 $ (mC $ s2 $ t),m)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   274
      | Const("HOL.divide",_) $ numt $ (Const("Numeral.number_of",_)$dent) =>
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   275
          let val den = HOLogic.dest_binum dent
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   276
          in if den = 0 then raise Zero
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   277
             else demult(mC $ numt $ t,ratmul(m, ratinv(rat_of_int den)))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   278
          end
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   279
      | _ => atomult(mC,s,t,m)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   280
      ) handle TERM _ => atomult(mC,s,t,m))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   281
  | demult(atom as Const("HOL.divide",_) $ t $ (Const("Numeral.number_of",_)$dent), m) =
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   282
      (let val den = HOLogic.dest_binum dent
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   283
       in if den = 0 then raise Zero else demult(t,ratmul(m, ratinv(rat_of_int den))) end
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   284
       handle TERM _ => (Some atom,m))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   285
  | demult(t as Const("Numeral.number_of",_)$n,m) =
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   286
      ((None,ratmul(m,rat_of_int(HOLogic.dest_binum n)))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   287
       handle TERM _ => (Some t,m))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   288
  | demult(atom,m) = (Some atom,m)
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   289
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   290
and atomult(mC,atom,t,m) = (case demult(t,m) of (None,m') => (Some atom,m')
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   291
                            | (Some t',m') => (Some(mC $ atom $ t'),m'))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   292
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   293
fun decomp2 inj_consts (rel,lhs,rhs) =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   294
let
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   295
(* Turn term into list of summand * multiplicity plus a constant *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   296
fun poly(Const("op +",_) $ s $ t, m, pi) = poly(s,m,poly(t,m,pi))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   297
  | poly(all as Const("op -",T) $ s $ t, m, pi) =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   298
      if nT T then add_atom(all,m,pi)
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   299
      else poly(s,m,poly(t,ratneg m,pi))
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   300
  | poly(Const("uminus",_) $ t, m, pi) = poly(t,ratneg m,pi)
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   301
  | poly(Const("0",_), _, pi) = pi
11464
ddea204de5bc turned translation for 1::nat into def.
nipkow
parents: 11334
diff changeset
   302
  | poly(Const("1",_), m, (p,i)) = (p,ratadd(i,m))
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   303
  | poly(Const("Suc",_)$t, m, (p,i)) = poly(t, m, (p,ratadd(i,m)))
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   304
  | poly(t as Const("op *",_) $ _ $ _, m, pi as (p,i)) =
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   305
      (case demult(t,m) of
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   306
         (None,m') => (p,ratadd(i,m))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   307
       | (Some u,m') => add_atom(u,m',pi))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   308
  | poly(t as Const("HOL.divide",_) $ _ $ _, m, pi as (p,i)) =
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   309
      (case demult(t,m) of
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   310
         (None,m') => (p,ratadd(i,m))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   311
       | (Some u,m') => add_atom(u,m',pi))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   312
  | poly(all as (Const("Numeral.number_of",_)$t,m,(p,i))) =
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   313
      ((p,ratadd(i,ratmul(m,rat_of_int(HOLogic.dest_binum t))))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   314
       handle TERM _ => add_atom all)
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   315
  | poly(all as Const f $ x, m, pi) =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   316
      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
   317
  | poly x  = add_atom x;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   318
10718
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   319
val (p,i) = poly(lhs,rat_of_int 1,([],rat_of_int 0))
c058f78c3544 rational arithmetic
nipkow
parents: 10693
diff changeset
   320
and (q,j) = poly(rhs,rat_of_int 1,([],rat_of_int 0))
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   321
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   322
  in case rel of
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   323
       "op <"  => Some(p,i,"<",q,j)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   324
     | "op <=" => Some(p,i,"<=",q,j)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   325
     | "op ="  => Some(p,i,"=",q,j)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   326
     | _       => None
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   327
  end handle Zero => None;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   328
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   329
fun negate(Some(x,i,rel,y,j,d)) = Some(x,i,"~"^rel,y,j,d)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   330
  | negate None = None;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   331
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   332
fun decomp1 (discrete,inj_consts) (T,xxx) =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   333
  (case T of
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   334
     Type("fun",[Type(D,[]),_]) =>
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   335
       (case assoc(discrete,D) of
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   336
          None => None
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   337
        | Some d => (case decomp2 inj_consts xxx of
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   338
                       None => None
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   339
                     | Some(p,i,rel,q,j) => Some(p,i,rel,q,j,d)))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   340
   | _ => None);
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   341
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   342
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
   343
  | decomp2 data (_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   344
      negate(decomp1 data (T,(rel,lhs,rhs)))
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   345
  | decomp2 data _ = None
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   346
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   347
fun decomp sg =
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   348
  let val {discrete, inj_consts, ...} = ArithTheoryData.get_sg sg
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   349
  in decomp2 (discrete,inj_consts) end
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   350
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   351
fun number_of(n,T) = HOLogic.number_of_const T $ (HOLogic.mk_bin n)
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   352
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   353
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   354
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   355
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   356
structure Fast_Arith =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   357
  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
   358
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   359
val fast_arith_tac = Fast_Arith.lin_arith_tac
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   360
and trace_arith    = Fast_Arith.trace;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   361
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   362
local
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   363
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   364
(* reduce contradictory <= to False.
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   365
   Most of the work is done by the cancel tactics.
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   366
*)
11464
ddea204de5bc turned translation for 1::nat into def.
nipkow
parents: 11334
diff changeset
   367
val add_rules = [add_0,add_0_right,Zero_not_Suc,Suc_not_Zero,le_0_eq,One_def];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   368
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   369
val add_mono_thms_nat = map (fn s => prove_goal (the_context ()) s
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   370
 (fn prems => [cut_facts_tac prems 1,
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   371
               blast_tac (claset() addIs [add_le_mono]) 1]))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   372
["(i <= j) & (k <= l) ==> i + k <= j + (l::nat)",
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   373
 "(i  = j) & (k <= l) ==> i + k <= j + (l::nat)",
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   374
 "(i <= j) & (k  = l) ==> i + k <= j + (l::nat)",
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   375
 "(i  = j) & (k  = l) ==> i + k  = j + (l::nat)"
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   376
];
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   377
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   378
in
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   379
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   380
val init_lin_arith_data =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   381
 Fast_Arith.setup @
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   382
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset = _} =>
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   383
   {add_mono_thms = add_mono_thms @ add_mono_thms_nat,
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   384
    mult_mono_thms = mult_mono_thms,
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   385
    inj_thms = inj_thms,
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   386
    lessD = lessD @ [Suc_leI],
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   387
    simpset = HOL_basic_ss addsimps add_rules addsimprocs nat_cancel_sums_add}),
9593
b732997cfc11 tuned names;
wenzelm
parents: 9436
diff changeset
   388
  ArithTheoryData.init, arith_discrete ("nat", true)];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   389
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   390
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   391
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   392
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   393
local
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   394
val nat_arith_simproc_pats =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   395
  map (fn s => Thm.read_cterm (Theory.sign_of (the_context ())) (s, HOLogic.boolT))
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   396
      ["(m::nat) < n","(m::nat) <= n", "(m::nat) = n"];
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   397
in
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   398
val fast_nat_arith_simproc = mk_simproc
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   399
  "fast_nat_arith" nat_arith_simproc_pats Fast_Arith.lin_arith_prover;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   400
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   401
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   402
(* 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
   403
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
   404
*not* themselves (in)equalities, because the latter activate
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   405
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
   406
solver all the time rather than add the additional check. *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   407
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   408
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   409
(* arith proof method *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   410
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   411
(* 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
   412
   in case there are lots of irrelevant terms involved;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   413
   elimination of min/max can be optimized:
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   414
   (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
   415
   (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
   416
*)
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   417
local
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   418
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   419
fun raw_arith_tac i st =
9593
b732997cfc11 tuned names;
wenzelm
parents: 9436
diff changeset
   420
  refute_tac (K true) (REPEAT o split_tac (#splits (ArithTheoryData.get_sg (Thm.sign_of_thm st))))
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   421
             ((REPEAT_DETERM o etac linorder_neqE) THEN' fast_arith_tac) i st;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   422
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   423
in
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   424
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents: 10516
diff changeset
   425
val arith_tac = fast_arith_tac ORELSE' (atomize_tac THEN' raw_arith_tac);
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   426
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   427
fun arith_method prems =
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   428
  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
   429
10516
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   430
end;
dc113303d101 arith_tac: atomize;
wenzelm
parents: 9893
diff changeset
   431
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   432
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   433
(* theory setup *)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   434
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   435
val arith_setup =
10766
ace2ba2d4fd1 removal of the nat_cancel_factor simproc
paulson
parents: 10718
diff changeset
   436
 [Simplifier.change_simpset_of (op addsimprocs) nat_cancel_sums] @
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   437
  init_lin_arith_data @
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   438
  [Simplifier.change_simpset_of (op addSolver)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   439
   (mk_solver "lin. arith." Fast_Arith.cut_lin_arith_tac),
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   440
  Simplifier.change_simpset_of (op addsimprocs) [fast_nat_arith_simproc],
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   441
  Method.add_methods [("arith", (arith_method o #2) oo Method.syntax Args.bang_facts,
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   442
    "decide linear arithmethic")],
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   443
  Attrib.add_attributes [("arith_split",
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   444
    (Attrib.no_args arith_split_add, Attrib.no_args Attrib.undef_local_attribute),
9893
93d2fde0306c tuned msg;
wenzelm
parents: 9593
diff changeset
   445
    "declaration of split rules for arithmetic procedure")]];