src/HOL/Real/rat_arith.ML
author paulson
Tue, 27 Jan 2004 15:39:51 +0100
changeset 14365 3d4df8c166ae
child 14368 2763da611ad9
permissions -rw-r--r--
replacing HOL/Real/PRat, PNat by the rational number development of Markus Wenzel
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Real/rat_arith0.ML
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     2
    ID:         $Id$
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     4
    Copyright   2004 University of Cambridge
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     5
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     6
Simprocs for common factor cancellation & Rational coefficient handling
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     7
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     8
Instantiation of the generic linear arithmetic package for type rat.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
     9
*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    10
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    11
(*FIXME DELETE*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    12
val rat_mult_strict_left_mono =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    13
    read_instantiate_sg(sign_of (the_context())) [("a","?a::rat")] mult_strict_left_mono;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    14
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    15
val rat_mult_left_mono =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    16
    read_instantiate_sg(sign_of (the_context())) [("a","?a::rat")] mult_left_mono;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    17
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    18
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    19
val rat_number_of_def = thm "rat_number_of_def";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    20
val diff_rat_def = thm "diff_rat_def";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    21
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    22
val rat_of_int_zero = thm "rat_of_int_zero";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    23
val rat_of_int_one = thm "rat_of_int_one";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    24
val rat_of_int_add_distrib = thm "rat_of_int_add_distrib";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    25
val rat_of_int_minus_distrib = thm "rat_of_int_minus_distrib";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    26
val rat_of_int_diff_distrib = thm "rat_of_int_diff_distrib";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    27
val rat_of_int_mult_distrib = thm "rat_of_int_mult_distrib";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    28
val rat_inject = thm "rat_inject";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    29
val rat_of_int_zero_cancel = thm "rat_of_int_zero_cancel";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    30
val rat_of_int_less_iff = thm "rat_of_int_less_iff";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    31
val rat_of_int_le_iff = thm "rat_of_int_le_iff";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    32
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    33
val number_of_rat = thm "number_of_rat";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    34
val rat_numeral_0_eq_0 = thm "rat_numeral_0_eq_0";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    35
val rat_numeral_1_eq_1 = thm "rat_numeral_1_eq_1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    36
val add_rat_number_of = thm "add_rat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    37
val minus_rat_number_of = thm "minus_rat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    38
val diff_rat_number_of = thm "diff_rat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    39
val mult_rat_number_of = thm "mult_rat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    40
val rat_mult_2 = thm "rat_mult_2";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    41
val rat_mult_2_right = thm "rat_mult_2_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    42
val eq_rat_number_of = thm "eq_rat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    43
val less_rat_number_of = thm "less_rat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    44
val rat_minus_1_eq_m1 = thm "rat_minus_1_eq_m1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    45
val rat_mult_minus1 = thm "rat_mult_minus1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    46
val rat_mult_minus1_right = thm "rat_mult_minus1_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    47
val rat_add_number_of_left = thm "rat_add_number_of_left";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    48
val rat_mult_number_of_left = thm "rat_mult_number_of_left";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    49
val rat_add_number_of_diff1 = thm "rat_add_number_of_diff1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    50
val rat_add_number_of_diff2 = thm "rat_add_number_of_diff2";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    51
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    52
val rat_add_0_left = thm "rat_add_0_left";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    53
val rat_add_0_right = thm "rat_add_0_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    54
val rat_mult_1_left = thm "rat_mult_1_left";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    55
val rat_mult_1_right = thm "rat_mult_1_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    56
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    57
(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    58
val rat_numeral_ss =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    59
    HOL_ss addsimps [rat_numeral_0_eq_0 RS sym, rat_numeral_1_eq_1 RS sym,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    60
                     rat_minus_1_eq_m1];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    61
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    62
fun rename_numerals th =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    63
    asm_full_simplify rat_numeral_ss (Thm.transfer (the_context ()) th);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    64
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    65
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    66
structure Rat_Numeral_Simprocs =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    67
struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    68
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    69
(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    70
  isn't complicated by the abstract 0 and 1.*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    71
val numeral_syms = [rat_numeral_0_eq_0 RS sym, rat_numeral_1_eq_1 RS sym];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    72
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    73
(*Utilities*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    74
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    75
val ratT = Type("Rational.rat", []);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    76
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    77
fun mk_numeral n = HOLogic.number_of_const ratT $ HOLogic.mk_bin n;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    78
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    79
(*Decodes a binary rat constant, or 0, 1*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    80
val dest_numeral = Int_Numeral_Simprocs.dest_numeral;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    81
val find_first_numeral = Int_Numeral_Simprocs.find_first_numeral;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    82
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    83
val zero = mk_numeral 0;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    84
val mk_plus = HOLogic.mk_binop "op +";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    85
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    86
val uminus_const = Const ("uminus", ratT --> ratT);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    87
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    88
(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    89
fun mk_sum []        = zero
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    90
  | mk_sum [t,u]     = mk_plus (t, u)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    91
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    92
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    93
(*this version ALWAYS includes a trailing zero*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    94
fun long_mk_sum []        = zero
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    95
  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    96
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    97
val dest_plus = HOLogic.dest_bin "op +" ratT;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    98
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
    99
(*decompose additions AND subtractions as a sum*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   100
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   101
        dest_summing (pos, t, dest_summing (pos, u, ts))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   102
  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   103
        dest_summing (pos, t, dest_summing (not pos, u, ts))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   104
  | dest_summing (pos, t, ts) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   105
        if pos then t::ts else uminus_const$t :: ts;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   106
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   107
fun dest_sum t = dest_summing (true, t, []);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   108
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   109
val mk_diff = HOLogic.mk_binop "op -";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   110
val dest_diff = HOLogic.dest_bin "op -" ratT;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   111
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   112
val one = mk_numeral 1;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   113
val mk_times = HOLogic.mk_binop "op *";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   114
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   115
fun mk_prod [] = one
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   116
  | mk_prod [t] = t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   117
  | mk_prod (t :: ts) = if t = one then mk_prod ts
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   118
                        else mk_times (t, mk_prod ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   119
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   120
val dest_times = HOLogic.dest_bin "op *" ratT;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   121
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   122
fun dest_prod t =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   123
      let val (t,u) = dest_times t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   124
      in  dest_prod t @ dest_prod u  end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   125
      handle TERM _ => [t];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   126
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   127
(*DON'T do the obvious simplifications; that would create special cases*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   128
fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   129
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   130
(*Express t as a product of (possibly) a numeral with other sorted terms*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   131
fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   132
  | dest_coeff sign t =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   133
    let val ts = sort Term.term_ord (dest_prod t)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   134
        val (n, ts') = find_first_numeral [] ts
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   135
                          handle TERM _ => (1, ts)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   136
    in (sign*n, mk_prod ts') end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   137
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   138
(*Find first coefficient-term THAT MATCHES u*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   139
fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   140
  | find_first_coeff past u (t::terms) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   141
        let val (n,u') = dest_coeff 1 t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   142
        in  if u aconv u' then (n, rev past @ terms)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   143
                          else find_first_coeff (t::past) u terms
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   144
        end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   145
        handle TERM _ => find_first_coeff (t::past) u terms;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   146
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   147
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   148
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   149
val add_0s  = map rename_numerals [rat_add_0_left, rat_add_0_right];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   150
val mult_1s = map rename_numerals [rat_mult_1_left, rat_mult_1_right] @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   151
              [rat_mult_minus1, rat_mult_minus1_right];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   152
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   153
(*To perform binary arithmetic*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   154
val bin_simps =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   155
    [rat_numeral_0_eq_0 RS sym, rat_numeral_1_eq_1 RS sym,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   156
     add_rat_number_of, rat_add_number_of_left, minus_rat_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   157
     diff_rat_number_of, mult_rat_number_of, rat_mult_number_of_left] @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   158
    bin_arith_simps @ bin_rel_simps;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   159
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   160
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   161
  during re-arrangement*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   162
val non_add_bin_simps = 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   163
    bin_simps \\ [rat_add_number_of_left, add_rat_number_of];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   164
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   165
(*To evaluate binary negations of coefficients*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   166
val rat_minus_simps = NCons_simps @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   167
                   [rat_minus_1_eq_m1, minus_rat_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   168
                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   169
                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   170
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   171
(*To let us treat subtraction as addition*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   172
val diff_simps = [diff_rat_def, minus_add_distrib, minus_minus];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   173
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   174
(*to extract again any uncancelled minuses*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   175
val rat_minus_from_mult_simps =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   176
    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   177
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   178
(*combine unary minus with numeric literals, however nested within a product*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   179
val rat_mult_minus_simps =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   180
    [mult_assoc, minus_mult_left, minus_mult_commute];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   181
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   182
(*Apply the given rewrite (if present) just once*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   183
fun trans_tac None      = all_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   184
  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   185
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   186
(*Final simplification: cancel + and *  *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   187
val simplify_meta_eq =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   188
    Int_Numeral_Simprocs.simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   189
         [add_0, add_0_right,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   190
          mult_zero_left, mult_zero_right, mult_1, mult_1_right];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   191
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   192
fun prep_simproc (name, pats, proc) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   193
  Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   194
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   195
structure CancelNumeralsCommon =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   196
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   197
  val mk_sum            = mk_sum
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   198
  val dest_sum          = dest_sum
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   199
  val mk_coeff          = mk_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   200
  val dest_coeff        = dest_coeff 1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   201
  val find_first_coeff  = find_first_coeff []
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   202
  val trans_tac         = trans_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   203
  val norm_tac =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   204
     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   205
                                         rat_minus_simps@add_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   206
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@rat_mult_minus_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   207
     THEN ALLGOALS
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   208
              (simp_tac (HOL_ss addsimps rat_minus_from_mult_simps@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   209
                                         add_ac@mult_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   210
  val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   211
  val simplify_meta_eq  = simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   212
  end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   213
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   214
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   215
structure EqCancelNumerals = CancelNumeralsFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   216
 (open CancelNumeralsCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   217
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   218
  val mk_bal   = HOLogic.mk_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   219
  val dest_bal = HOLogic.dest_bin "op =" ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   220
  val bal_add1 = eq_add_iff1 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   221
  val bal_add2 = eq_add_iff2 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   222
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   223
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   224
structure LessCancelNumerals = CancelNumeralsFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   225
 (open CancelNumeralsCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   226
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   227
  val mk_bal   = HOLogic.mk_binrel "op <"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   228
  val dest_bal = HOLogic.dest_bin "op <" ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   229
  val bal_add1 = less_add_iff1 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   230
  val bal_add2 = less_add_iff2 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   231
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   232
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   233
structure LeCancelNumerals = CancelNumeralsFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   234
 (open CancelNumeralsCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   235
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   236
  val mk_bal   = HOLogic.mk_binrel "op <="
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   237
  val dest_bal = HOLogic.dest_bin "op <=" ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   238
  val bal_add1 = le_add_iff1 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   239
  val bal_add2 = le_add_iff2 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   240
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   241
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   242
val cancel_numerals =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   243
  map prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   244
   [("rateq_cancel_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   245
     ["(l::rat) + m = n", "(l::rat) = m + n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   246
      "(l::rat) - m = n", "(l::rat) = m - n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   247
      "(l::rat) * m = n", "(l::rat) = m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   248
     EqCancelNumerals.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   249
    ("ratless_cancel_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   250
     ["(l::rat) + m < n", "(l::rat) < m + n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   251
      "(l::rat) - m < n", "(l::rat) < m - n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   252
      "(l::rat) * m < n", "(l::rat) < m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   253
     LessCancelNumerals.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   254
    ("ratle_cancel_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   255
     ["(l::rat) + m <= n", "(l::rat) <= m + n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   256
      "(l::rat) - m <= n", "(l::rat) <= m - n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   257
      "(l::rat) * m <= n", "(l::rat) <= m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   258
     LeCancelNumerals.proc)];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   259
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   260
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   261
structure CombineNumeralsData =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   262
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   263
  val add               = op + : int*int -> int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   264
  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   265
  val dest_sum          = dest_sum
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   266
  val mk_coeff          = mk_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   267
  val dest_coeff        = dest_coeff 1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   268
  val left_distrib      = combine_common_factor RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   269
  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   270
  val trans_tac         = trans_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   271
  val norm_tac =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   272
     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   273
                                   diff_simps@rat_minus_simps@add_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   274
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@rat_mult_minus_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   275
     THEN ALLGOALS (simp_tac (HOL_ss addsimps rat_minus_from_mult_simps@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   276
                                              add_ac@mult_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   277
  val numeral_simp_tac  = ALLGOALS
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   278
                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   279
  val simplify_meta_eq  =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   280
        Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   281
  end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   282
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   283
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   284
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   285
val combine_numerals =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   286
  prep_simproc ("rat_combine_numerals", ["(i::rat) + j", "(i::rat) - j"], CombineNumerals.proc);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   287
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   288
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   289
(** Declarations for ExtractCommonTerm **)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   290
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   291
(*this version ALWAYS includes a trailing one*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   292
fun long_mk_prod []        = one
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   293
  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   294
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   295
(*Find first term that matches u*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   296
fun find_first past u []         = raise TERM("find_first", [])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   297
  | find_first past u (t::terms) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   298
        if u aconv t then (rev past @ terms)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   299
        else find_first (t::past) u terms
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   300
        handle TERM _ => find_first (t::past) u terms;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   301
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   302
(*Final simplification: cancel + and *  *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   303
fun cancel_simplify_meta_eq cancel_th th =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   304
    Int_Numeral_Simprocs.simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   305
        [rat_mult_1_left, rat_mult_1_right]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   306
        (([th, cancel_th]) MRS trans);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   307
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   308
(*** Making constant folding work for 0 and 1 too ***)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   309
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   310
structure RatAbstractNumeralsData =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   311
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   312
  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   313
  val is_numeral      = Bin_Simprocs.is_numeral
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   314
  val numeral_0_eq_0  = rat_numeral_0_eq_0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   315
  val numeral_1_eq_1  = rat_numeral_1_eq_1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   316
  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   317
  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   318
  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   319
  end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   320
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   321
structure RatAbstractNumerals = AbstractNumeralsFun (RatAbstractNumeralsData)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   322
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   323
(*For addition, we already have rules for the operand 0.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   324
  Multiplication is omitted because there are already special rules for
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   325
  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   326
  For the others, having three patterns is a compromise between just having
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   327
  one (many spurious calls) and having nine (just too many!) *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   328
val eval_numerals =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   329
  map prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   330
   [("rat_add_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   331
     ["(m::rat) + 1", "(m::rat) + number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   332
     RatAbstractNumerals.proc add_rat_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   333
    ("rat_diff_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   334
     ["(m::rat) - 1", "(m::rat) - number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   335
     RatAbstractNumerals.proc diff_rat_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   336
    ("rat_eq_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   337
     ["(m::rat) = 0", "(m::rat) = 1", "(m::rat) = number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   338
     RatAbstractNumerals.proc eq_rat_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   339
    ("rat_less_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   340
     ["(m::rat) < 0", "(m::rat) < 1", "(m::rat) < number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   341
     RatAbstractNumerals.proc less_rat_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   342
    ("rat_le_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   343
     ["(m::rat) <= 0", "(m::rat) <= 1", "(m::rat) <= number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   344
     RatAbstractNumerals.proc le_number_of_eq_not_less)]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   345
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   346
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   347
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   348
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   349
Addsimprocs Rat_Numeral_Simprocs.eval_numerals;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   350
Addsimprocs Rat_Numeral_Simprocs.cancel_numerals;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   351
Addsimprocs [Rat_Numeral_Simprocs.combine_numerals];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   352
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   353
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   354
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   355
(** Constant folding for rat plus and times **)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   356
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   357
(*We do not need
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   358
    structure Rat_Plus_Assoc = Assoc_Fold (Rat_Plus_Assoc_Data);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   359
  because combine_numerals does the same thing*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   360
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   361
structure Rat_Times_Assoc_Data : ASSOC_FOLD_DATA =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   362
struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   363
  val ss                = HOL_ss
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   364
  val eq_reflection     = eq_reflection
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   365
  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   366
  val T      = Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   367
  val plus   = Const ("op *", [Rat_Numeral_Simprocs.ratT,Rat_Numeral_Simprocs.ratT] ---> Rat_Numeral_Simprocs.ratT)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   368
  val add_ac = mult_ac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   369
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   370
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   371
structure Rat_Times_Assoc = Assoc_Fold (Rat_Times_Assoc_Data);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   372
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   373
Addsimprocs [Rat_Times_Assoc.conv];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   374
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   375
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   376
(****Common factor cancellation****)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   377
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   378
(*To quote from Provers/Arith/cancel_numeral_factor.ML:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   379
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   380
This simproc Cancels common coefficients in balanced expressions:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   381
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   382
     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   383
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   384
where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   385
and d = gcd(m,m') and n=m/d and n'=m'/d.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   386
*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   387
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   388
local
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   389
  open Rat_Numeral_Simprocs
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   390
in
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   391
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   392
val rel_rat_number_of = [eq_rat_number_of, less_rat_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   393
                          le_number_of_eq_not_less]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   394
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   395
structure CancelNumeralFactorCommon =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   396
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   397
  val mk_coeff          = mk_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   398
  val dest_coeff        = dest_coeff 1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   399
  val trans_tac         = trans_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   400
  val norm_tac =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   401
     ALLGOALS (simp_tac (HOL_ss addsimps rat_minus_from_mult_simps @ mult_1s))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   402
     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@rat_mult_minus_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   403
     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   404
  val numeral_simp_tac  =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   405
         ALLGOALS (simp_tac (HOL_ss addsimps rel_rat_number_of@bin_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   406
  val simplify_meta_eq  = simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   407
  end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   408
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   409
structure DivCancelNumeralFactor = CancelNumeralFactorFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   410
 (open CancelNumeralFactorCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   411
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   412
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   413
  val dest_bal = HOLogic.dest_bin "HOL.divide" Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   414
  val cancel = mult_divide_cancel_left RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   415
  val neg_exchanges = false
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   416
)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   417
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   418
structure EqCancelNumeralFactor = CancelNumeralFactorFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   419
 (open CancelNumeralFactorCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   420
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   421
  val mk_bal   = HOLogic.mk_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   422
  val dest_bal = HOLogic.dest_bin "op =" Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   423
  val cancel = mult_cancel_left RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   424
  val neg_exchanges = false
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   425
)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   426
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   427
structure LessCancelNumeralFactor = CancelNumeralFactorFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   428
 (open CancelNumeralFactorCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   429
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   430
  val mk_bal   = HOLogic.mk_binrel "op <"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   431
  val dest_bal = HOLogic.dest_bin "op <" Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   432
  val cancel = mult_less_cancel_left RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   433
  val neg_exchanges = true
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   434
)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   435
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   436
structure LeCancelNumeralFactor = CancelNumeralFactorFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   437
 (open CancelNumeralFactorCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   438
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   439
  val mk_bal   = HOLogic.mk_binrel "op <="
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   440
  val dest_bal = HOLogic.dest_bin "op <=" Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   441
  val cancel = mult_le_cancel_left RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   442
  val neg_exchanges = true
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   443
)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   444
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   445
val rat_cancel_numeral_factors_relations =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   446
  map prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   447
   [("rateq_cancel_numeral_factor",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   448
     ["(l::rat) * m = n", "(l::rat) = m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   449
     EqCancelNumeralFactor.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   450
    ("ratless_cancel_numeral_factor",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   451
     ["(l::rat) * m < n", "(l::rat) < m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   452
     LessCancelNumeralFactor.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   453
    ("ratle_cancel_numeral_factor",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   454
     ["(l::rat) * m <= n", "(l::rat) <= m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   455
     LeCancelNumeralFactor.proc)]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   456
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   457
val rat_cancel_numeral_factors_divide = prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   458
        ("ratdiv_cancel_numeral_factor",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   459
         ["((l::rat) * m) / n", "(l::rat) / (m * n)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   460
          "((number_of v)::rat) / (number_of w)"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   461
         DivCancelNumeralFactor.proc)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   462
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   463
val rat_cancel_numeral_factors =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   464
    rat_cancel_numeral_factors_relations @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   465
    [rat_cancel_numeral_factors_divide]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   466
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   467
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   468
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   469
Addsimprocs rat_cancel_numeral_factors;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   470
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   471
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   472
(*examples:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   473
print_depth 22;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   474
set timing;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   475
set trace_simp;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   476
fun test s = (Goal s; by (Simp_tac 1));
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   477
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   478
test "0 <= (y::rat) * -2";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   479
test "9*x = 12 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   480
test "(9*x) / (12 * (y::rat)) = z";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   481
test "9*x < 12 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   482
test "9*x <= 12 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   483
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   484
test "-99*x = 132 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   485
test "(-99*x) / (132 * (y::rat)) = z";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   486
test "-99*x < 132 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   487
test "-99*x <= 132 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   488
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   489
test "999*x = -396 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   490
test "(999*x) / (-396 * (y::rat)) = z";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   491
test "999*x < -396 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   492
test "999*x <= -396 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   493
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   494
test  "(- ((2::rat) * x) <= 2 * y)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   495
test "-99*x = -81 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   496
test "(-99*x) / (-81 * (y::rat)) = z";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   497
test "-99*x <= -81 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   498
test "-99*x < -81 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   499
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   500
test "-2 * x = -1 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   501
test "-2 * x = -(y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   502
test "(-2 * x) / (-1 * (y::rat)) = z";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   503
test "-2 * x < -(y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   504
test "-2 * x <= -1 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   505
test "-x < -23 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   506
test "-x <= -23 * (y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   507
*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   508
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   509
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   510
(** Declarations for ExtractCommonTerm **)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   511
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   512
local
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   513
  open Rat_Numeral_Simprocs
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   514
in
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   515
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   516
structure CancelFactorCommon =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   517
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   518
  val mk_sum            = long_mk_prod
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   519
  val dest_sum          = dest_prod
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   520
  val mk_coeff          = mk_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   521
  val dest_coeff        = dest_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   522
  val find_first        = find_first []
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   523
  val trans_tac         = trans_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   524
  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   525
  end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   526
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   527
structure EqCancelFactor = ExtractCommonTermFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   528
 (open CancelFactorCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   529
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   530
  val mk_bal   = HOLogic.mk_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   531
  val dest_bal = HOLogic.dest_bin "op =" Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   532
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_cancel_left
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   533
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   534
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   535
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   536
structure DivideCancelFactor = ExtractCommonTermFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   537
 (open CancelFactorCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   538
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   539
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   540
  val dest_bal = HOLogic.dest_bin "HOL.divide" Rat_Numeral_Simprocs.ratT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   541
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   542
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   543
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   544
val rat_cancel_factor =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   545
  map prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   546
   [("rat_eq_cancel_factor", ["(l::rat) * m = n", "(l::rat) = m * n"], EqCancelFactor.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   547
    ("rat_divide_cancel_factor", ["((l::rat) * m) / n", "(l::rat) / (m * n)"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   548
     DivideCancelFactor.proc)];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   549
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   550
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   551
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   552
Addsimprocs rat_cancel_factor;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   553
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   554
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   555
(*examples:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   556
print_depth 22;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   557
set timing;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   558
set trace_simp;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   559
fun test s = (Goal s; by (Asm_simp_tac 1));
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   560
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   561
test "x*k = k*(y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   562
test "k = k*(y::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   563
test "a*(b*c) = (b::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   564
test "a*(b*c) = d*(b::rat)*(x*a)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   565
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   566
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   567
test "(x*k) / (k*(y::rat)) = (uu::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   568
test "(k) / (k*(y::rat)) = (uu::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   569
test "(a*(b*c)) / ((b::rat)) = (uu::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   570
test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   571
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   572
(*FIXME: what do we do about this?*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   573
test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   574
*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   575
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   576
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   577
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   578
(****Instantiation of the generic linear arithmetic package****)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   579
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   580
val add_zero_left = thm"Ring_and_Field.add_0";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   581
val add_zero_right = thm"Ring_and_Field.add_0_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   582
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   583
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   584
local
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   585
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   586
(* reduce contradictory <= to False *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   587
val add_rules = 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   588
    [order_less_irrefl, rat_numeral_0_eq_0, rat_numeral_1_eq_1,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   589
     rat_minus_1_eq_m1, 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   590
     add_rat_number_of, minus_rat_number_of, diff_rat_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   591
     mult_rat_number_of, eq_rat_number_of, less_rat_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   592
     le_number_of_eq_not_less, diff_minus,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   593
     minus_add_distrib, minus_minus, mult_assoc, minus_zero,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   594
     add_zero_left, add_zero_right, left_minus, right_minus,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   595
     mult_zero_left, mult_zero_right, mult_1, mult_1_right,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   596
     minus_mult_left RS sym, minus_mult_right RS sym];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   597
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   598
val simprocs = [Rat_Times_Assoc.conv, Rat_Numeral_Simprocs.combine_numerals,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   599
                rat_cancel_numeral_factors_divide]@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   600
               Rat_Numeral_Simprocs.cancel_numerals @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   601
               Rat_Numeral_Simprocs.eval_numerals;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   602
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   603
val mono_ss = simpset() addsimps
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   604
                [add_mono,add_strict_mono,add_less_le_mono,add_le_less_mono];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   605
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   606
val add_mono_thms_rat =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   607
  map (fn s => prove_goal (the_context ()) s
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   608
                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   609
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   610
     "(i  = j) & (k <= l) ==> i + k <= j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   611
     "(i <= j) & (k  = l) ==> i + k <= j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   612
     "(i  = j) & (k  = l) ==> i + k  = j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   613
     "(i < j) & (k = l)   ==> i + k < j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   614
     "(i = j) & (k < l)   ==> i + k < j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   615
     "(i < j) & (k <= l)  ==> i + k < j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   616
     "(i <= j) & (k < l)  ==> i + k < j + (l::rat)",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   617
     "(i < j) & (k < l)   ==> i + k < j + (l::rat)"];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   618
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   619
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   620
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   621
val rat_mult_mono_thms =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   622
 [(rat_mult_strict_left_mono,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   623
   cvar(rat_mult_strict_left_mono, hd(tl(prems_of rat_mult_strict_left_mono)))),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   624
  (rat_mult_left_mono,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   625
   cvar(rat_mult_left_mono, hd(tl(prems_of rat_mult_left_mono))))]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   626
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   627
(* reduce contradictory <= to False *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   628
val simps = [True_implies_equals,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   629
             inst "a" "(number_of ?v)::rat" right_distrib,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   630
             divide_1,times_divide_eq_right,times_divide_eq_left,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   631
         rat_of_int_zero, rat_of_int_one, rat_of_int_add_distrib,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   632
         rat_of_int_minus_distrib, rat_of_int_diff_distrib,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   633
         rat_of_int_mult_distrib, number_of_rat RS sym];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   634
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   635
in
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   636
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   637
val fast_rat_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   638
  "fast_rat_arith" ["(m::rat) < n","(m::rat) <= n", "(m::rat) = n"]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   639
  Fast_Arith.lin_arith_prover;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   640
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   641
val int_inj_thms = [rat_of_int_le_iff RS iffD2, rat_of_int_less_iff RS iffD2,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   642
                    rat_inject RS iffD2];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   643
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   644
val rat_arith_setup =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   645
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   646
   {add_mono_thms = add_mono_thms @ add_mono_thms_rat,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   647
    mult_mono_thms = mult_mono_thms @ rat_mult_mono_thms,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   648
    inj_thms = (***int_inj_thms @*???**)  inj_thms,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   649
    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the rats are dense!*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   650
    simpset = simpset addsimps add_rules
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   651
                      addsimps simps
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   652
                      addsimprocs simprocs}),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   653
(*???
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   654
  arith_inj_const ("Rational.rat", HOLogic.intT --> Rat_Numeral_Simprocs.ratT),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   655
???*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   656
  arith_discrete ("Rational.rat",false),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   657
  Simplifier.change_simpset_of (op addsimprocs) [fast_rat_arith_simproc]];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   658
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   659
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   660
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   661
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
diff changeset
   662