src/HOL/Real/real_arith.ML
author paulson
Wed, 28 Jan 2004 17:01:01 +0100
changeset 14369 c50188fe6366
parent 14368 2763da611ad9
child 14387 e96d5c42c4b0
permissions -rw-r--r--
tidying up arithmetic for the hyperreals
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
     1
(*  Title:      HOL/Real/real_arith0.ML
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     3
    Author:     Tobias Nipkow, TU Muenchen
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
     4
    Copyright   1999 TU Muenchen
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     5
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
     6
Simprocs for common factor cancellation & Rational coefficient handling
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
     7
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
     8
Instantiation of the generic linear arithmetic package for type real.
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     9
*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    10
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    11
(*FIXME DELETE*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    12
val real_mult_left_mono =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    13
    read_instantiate_sg(sign_of (the_context())) [("a","?a::real")] mult_left_mono;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    14
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    15
val real_numeral_0_eq_0 = thm "real_numeral_0_eq_0";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    16
val real_numeral_1_eq_1 = thm "real_numeral_1_eq_1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    17
val real_number_of_def = thm "real_number_of_def";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    18
val add_real_number_of = thm "add_real_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    19
val minus_real_number_of = thm "minus_real_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    20
val diff_real_number_of = thm "diff_real_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    21
val mult_real_number_of = thm "mult_real_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    22
val real_mult_2 = thm "real_mult_2";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    23
val real_mult_2_right = thm "real_mult_2_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    24
val eq_real_number_of = thm "eq_real_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    25
val less_real_number_of = thm "less_real_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    26
val real_minus_1_eq_m1 = thm "real_minus_1_eq_m1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    27
val real_mult_minus1 = thm "real_mult_minus1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    28
val real_mult_minus1_right = thm "real_mult_minus1_right";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    29
val zero_less_real_of_nat_iff = thm "zero_less_real_of_nat_iff";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    30
val zero_le_real_of_nat_iff = thm "zero_le_real_of_nat_iff";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    31
val real_add_number_of_left = thm "real_add_number_of_left";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    32
val real_mult_number_of_left = thm "real_mult_number_of_left";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    33
val real_add_number_of_diff1 = thm "real_add_number_of_diff1";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    34
val real_add_number_of_diff2 = thm "real_add_number_of_diff2";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    35
val real_of_nat_number_of = thm "real_of_nat_number_of";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    36
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    37
(*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: 14356
diff changeset
    38
val real_numeral_ss =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    39
    HOL_ss addsimps [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    40
                     real_minus_1_eq_m1];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    41
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    42
fun rename_numerals th =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    43
    asm_full_simplify real_numeral_ss (Thm.transfer (the_context ()) th);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    44
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    45
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    46
structure Real_Numeral_Simprocs =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    47
struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    48
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    49
(*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: 14356
diff changeset
    50
  isn't complicated by the abstract 0 and 1.*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    51
val numeral_syms = [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    52
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    53
(*Utilities*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    54
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    55
fun mk_numeral n = HOLogic.number_of_const HOLogic.realT $ HOLogic.mk_bin n;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    56
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    57
(*Decodes a binary real constant, or 0, 1*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    58
val dest_numeral = Int_Numeral_Simprocs.dest_numeral;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    59
val find_first_numeral = Int_Numeral_Simprocs.find_first_numeral;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    60
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    61
val zero = mk_numeral 0;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    62
val mk_plus = HOLogic.mk_binop "op +";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    63
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    64
val uminus_const = Const ("uminus", HOLogic.realT --> HOLogic.realT);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    65
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    66
(*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: 14356
diff changeset
    67
fun mk_sum []        = zero
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    68
  | mk_sum [t,u]     = mk_plus (t, u)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    69
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    70
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    71
(*this version ALWAYS includes a trailing zero*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    72
fun long_mk_sum []        = zero
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    73
  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    74
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    75
val dest_plus = HOLogic.dest_bin "op +" HOLogic.realT;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    76
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    77
(*decompose additions AND subtractions as a sum*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    78
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    79
        dest_summing (pos, t, dest_summing (pos, u, ts))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    80
  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    81
        dest_summing (pos, t, dest_summing (not pos, u, ts))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    82
  | dest_summing (pos, t, ts) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    83
        if pos then t::ts else uminus_const$t :: ts;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    84
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    85
fun dest_sum t = dest_summing (true, t, []);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    86
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    87
val mk_diff = HOLogic.mk_binop "op -";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    88
val dest_diff = HOLogic.dest_bin "op -" HOLogic.realT;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    89
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    90
val one = mk_numeral 1;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    91
val mk_times = HOLogic.mk_binop "op *";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    92
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    93
fun mk_prod [] = one
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    94
  | mk_prod [t] = t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    95
  | mk_prod (t :: ts) = if t = one then mk_prod ts
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    96
                        else mk_times (t, mk_prod ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    97
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    98
val dest_times = HOLogic.dest_bin "op *" HOLogic.realT;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
    99
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   100
fun dest_prod t =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   101
      let val (t,u) = dest_times t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   102
      in  dest_prod t @ dest_prod u  end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   103
      handle TERM _ => [t];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   104
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   105
(*DON'T do the obvious simplifications; that would create special cases*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   106
fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   107
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   108
(*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: 14356
diff changeset
   109
fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   110
  | dest_coeff sign t =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   111
    let val ts = sort Term.term_ord (dest_prod t)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   112
        val (n, ts') = find_first_numeral [] ts
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   113
                          handle TERM _ => (1, ts)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   114
    in (sign*n, mk_prod ts') end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   115
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   116
(*Find first coefficient-term THAT MATCHES u*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   117
fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   118
  | find_first_coeff past u (t::terms) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   119
        let val (n,u') = dest_coeff 1 t
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   120
        in  if u aconv u' then (n, rev past @ terms)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   121
                          else find_first_coeff (t::past) u terms
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   122
        end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   123
        handle TERM _ => find_first_coeff (t::past) u terms;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   124
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   125
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   126
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   127
val add_0s  = map rename_numerals [real_add_zero_left, real_add_zero_right];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   128
val mult_1s = map rename_numerals [real_mult_1, real_mult_1_right] @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   129
              [real_mult_minus1, real_mult_minus1_right];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   130
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   131
(*To perform binary arithmetic*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   132
val bin_simps =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   133
    [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   134
     add_real_number_of, real_add_number_of_left, minus_real_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   135
     diff_real_number_of, mult_real_number_of, real_mult_number_of_left] @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   136
    bin_arith_simps @ bin_rel_simps;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   137
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   138
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   139
  during re-arrangement*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   140
val non_add_bin_simps = 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   141
    bin_simps \\ [real_add_number_of_left, add_real_number_of];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   142
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   143
(*To evaluate binary negations of coefficients*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   144
val real_minus_simps = NCons_simps @
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   145
                   [real_minus_1_eq_m1, minus_real_number_of,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   146
                    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: 14356
diff changeset
   147
                    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: 14356
diff changeset
   148
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   149
(*To let us treat subtraction as addition*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   150
val diff_simps = [real_diff_def, minus_add_distrib, minus_minus];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   151
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   152
(*to extract again any uncancelled minuses*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   153
val real_minus_from_mult_simps =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   154
    [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: 14356
diff changeset
   155
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   156
(*combine unary minus with numeric literals, however nested within a product*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   157
val real_mult_minus_simps =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   158
    [mult_assoc, minus_mult_left, minus_mult_commute];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   159
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   160
(*Apply the given rewrite (if present) just once*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   161
fun trans_tac None      = all_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   162
  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   163
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   164
(*Final simplification: cancel + and *  *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   165
val simplify_meta_eq =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   166
    Int_Numeral_Simprocs.simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   167
         [add_0, add_0_right,
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   168
          mult_zero_left, mult_zero_right, mult_1, mult_1_right];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   169
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   170
fun prep_simproc (name, pats, proc) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   171
  Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   172
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   173
structure CancelNumeralsCommon =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   174
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   175
  val mk_sum            = mk_sum
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   176
  val dest_sum          = dest_sum
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   177
  val mk_coeff          = mk_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   178
  val dest_coeff        = dest_coeff 1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   179
  val find_first_coeff  = find_first_coeff []
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   180
  val trans_tac         = trans_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   181
  val norm_tac =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   182
     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: 14356
diff changeset
   183
                                         real_minus_simps@add_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   184
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   185
     THEN ALLGOALS
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   186
              (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   187
                                         add_ac@mult_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   188
  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: 14356
diff changeset
   189
  val simplify_meta_eq  = simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   190
  end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   191
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   192
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   193
structure EqCancelNumerals = CancelNumeralsFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   194
 (open CancelNumeralsCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   195
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   196
  val mk_bal   = HOLogic.mk_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   197
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   198
  val bal_add1 = eq_add_iff1 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   199
  val bal_add2 = eq_add_iff2 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   200
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   201
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   202
structure LessCancelNumerals = CancelNumeralsFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   203
 (open CancelNumeralsCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   204
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   205
  val mk_bal   = HOLogic.mk_binrel "op <"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   206
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   207
  val bal_add1 = less_add_iff1 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   208
  val bal_add2 = less_add_iff2 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   209
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   210
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   211
structure LeCancelNumerals = CancelNumeralsFun
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   212
 (open CancelNumeralsCommon
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   213
  val prove_conv = Bin_Simprocs.prove_conv
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   214
  val mk_bal   = HOLogic.mk_binrel "op <="
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   215
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   216
  val bal_add1 = le_add_iff1 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   217
  val bal_add2 = le_add_iff2 RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   218
);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   219
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   220
val cancel_numerals =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   221
  map prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   222
   [("realeq_cancel_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   223
     ["(l::real) + m = n", "(l::real) = m + n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   224
      "(l::real) - m = n", "(l::real) = m - n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   225
      "(l::real) * m = n", "(l::real) = m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   226
     EqCancelNumerals.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   227
    ("realless_cancel_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   228
     ["(l::real) + m < n", "(l::real) < m + n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   229
      "(l::real) - m < n", "(l::real) < m - n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   230
      "(l::real) * m < n", "(l::real) < m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   231
     LessCancelNumerals.proc),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   232
    ("realle_cancel_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   233
     ["(l::real) + m <= n", "(l::real) <= m + n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   234
      "(l::real) - m <= n", "(l::real) <= m - n",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   235
      "(l::real) * m <= n", "(l::real) <= m * n"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   236
     LeCancelNumerals.proc)];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   237
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   238
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   239
structure CombineNumeralsData =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   240
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   241
  val add               = op + : int*int -> int
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   242
  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: 14356
diff changeset
   243
  val dest_sum          = dest_sum
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   244
  val mk_coeff          = mk_coeff
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   245
  val dest_coeff        = dest_coeff 1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   246
  val left_distrib      = combine_common_factor RS trans
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   247
  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   248
  val trans_tac         = trans_tac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   249
  val norm_tac =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   250
     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: 14356
diff changeset
   251
                                   diff_simps@real_minus_simps@add_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   252
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   253
     THEN ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   254
                                              add_ac@mult_ac))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   255
  val numeral_simp_tac  = ALLGOALS
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   256
                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   257
  val simplify_meta_eq  =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   258
        Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   259
  end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   260
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   261
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   262
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   263
val combine_numerals =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   264
  prep_simproc ("real_combine_numerals", ["(i::real) + j", "(i::real) - j"], CombineNumerals.proc);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   265
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   266
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   267
(** Declarations for ExtractCommonTerm **)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   268
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   269
(*this version ALWAYS includes a trailing one*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   270
fun long_mk_prod []        = one
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   271
  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   272
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   273
(*Find first term that matches u*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   274
fun find_first past u []         = raise TERM("find_first", [])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   275
  | find_first past u (t::terms) =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   276
        if u aconv t then (rev past @ terms)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   277
        else find_first (t::past) u terms
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   278
        handle TERM _ => find_first (t::past) u terms;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   279
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   280
(*Final simplification: cancel + and *  *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   281
fun cancel_simplify_meta_eq cancel_th th =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   282
    Int_Numeral_Simprocs.simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   283
        [real_mult_1, real_mult_1_right]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   284
        (([th, cancel_th]) MRS trans);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   285
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   286
(*** Making constant folding work for 0 and 1 too ***)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   287
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   288
structure RealAbstractNumeralsData =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   289
  struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   290
  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: 14356
diff changeset
   291
  val is_numeral      = Bin_Simprocs.is_numeral
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   292
  val numeral_0_eq_0  = real_numeral_0_eq_0
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   293
  val numeral_1_eq_1  = real_numeral_1_eq_1
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   294
  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   295
  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   296
  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   297
  end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   298
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   299
structure RealAbstractNumerals = AbstractNumeralsFun (RealAbstractNumeralsData)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   300
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   301
(*For addition, we already have rules for the operand 0.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   302
  Multiplication is omitted because there are already special rules for
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   303
  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: 14356
diff changeset
   304
  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: 14356
diff changeset
   305
  one (many spurious calls) and having nine (just too many!) *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   306
val eval_numerals =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   307
  map prep_simproc
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   308
   [("real_add_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   309
     ["(m::real) + 1", "(m::real) + number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   310
     RealAbstractNumerals.proc add_real_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   311
    ("real_diff_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   312
     ["(m::real) - 1", "(m::real) - number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   313
     RealAbstractNumerals.proc diff_real_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   314
    ("real_eq_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   315
     ["(m::real) = 0", "(m::real) = 1", "(m::real) = number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   316
     RealAbstractNumerals.proc eq_real_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   317
    ("real_less_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   318
     ["(m::real) < 0", "(m::real) < 1", "(m::real) < number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   319
     RealAbstractNumerals.proc less_real_number_of),
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   320
    ("real_le_eval_numerals",
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   321
     ["(m::real) <= 0", "(m::real) <= 1", "(m::real) <= number_of v"],
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   322
     RealAbstractNumerals.proc le_number_of_eq_not_less)]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   323
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   324
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   325
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   326
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   327
Addsimprocs Real_Numeral_Simprocs.eval_numerals;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   328
Addsimprocs Real_Numeral_Simprocs.cancel_numerals;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   329
Addsimprocs [Real_Numeral_Simprocs.combine_numerals];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   330
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   331
(*examples:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   332
print_depth 22;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   333
set timing;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   334
set trace_simp;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   335
fun test s = (Goal s; by (Simp_tac 1));
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   336
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   337
test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   338
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   339
test "2*u = (u::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   340
test "(i + j + 12 + (k::real)) - 15 = y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   341
test "(i + j + 12 + (k::real)) - 5 = y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   342
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   343
test "y - b < (b::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   344
test "y - (3*b + c) < (b::real) - 2*c";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   345
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   346
test "(2*x - (u*v) + y) - v*3*u = (w::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   347
test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   348
test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   349
test "u*v - (x*u*v + (u*v)*4 + y) = (w::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   350
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   351
test "(i + j + 12 + (k::real)) = u + 15 + y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   352
test "(i + j*2 + 12 + (k::real)) = j + 5 + y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   353
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   354
test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   355
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   356
test "a + -(b+c) + b = (d::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   357
test "a + -(b+c) - b = (d::real)";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   358
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   359
(*negative numerals*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   360
test "(i + j + -2 + (k::real)) - (u + 5 + y) = zz";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   361
test "(i + j + -3 + (k::real)) < u + 5 + y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   362
test "(i + j + 3 + (k::real)) < u + -6 + y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   363
test "(i + j + -12 + (k::real)) - 15 = y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   364
test "(i + j + 12 + (k::real)) - -15 = y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   365
test "(i + j + -12 + (k::real)) - -15 = y";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   366
*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   367
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   368
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   369
(** Constant folding for real plus and times **)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   370
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   371
(*We do not need
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   372
    structure Real_Plus_Assoc = Assoc_Fold (Real_Plus_Assoc_Data);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   373
  because combine_numerals does the same thing*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   374
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   375
structure Real_Times_Assoc_Data : ASSOC_FOLD_DATA =
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   376
struct
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   377
  val ss                = HOL_ss
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   378
  val eq_reflection     = eq_reflection
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   379
  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   380
  val T      = HOLogic.realT
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   381
  val plus   = Const ("op *", [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   382
  val add_ac = mult_ac
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   383
end;
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   384
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   385
structure Real_Times_Assoc = Assoc_Fold (Real_Times_Assoc_Data);
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   386
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   387
Addsimprocs [Real_Times_Assoc.conv];
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   388
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   389
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   390
(****Common factor cancellation****)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   391
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   392
(*To quote from Provers/Arith/cancel_numeral_factor.ML:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   393
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   394
This simproc Cancels common coefficients in balanced expressions:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   395
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   396
     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   397
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   398
where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   399
and d = gcd(m,m') and n=m/d and n'=m'/d.
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   400
*)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
   401
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   402
local
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   403
  open Real_Numeral_Simprocs
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   404
in
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   405
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   406
val rel_real_number_of = [eq_real_number_of, less_real_number_of,
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   407
                          le_number_of_eq_not_less]
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   408
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   409
structure CancelNumeralFactorCommon =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   410
  struct
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   411
  val mk_coeff          = mk_coeff
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   412
  val dest_coeff        = dest_coeff 1
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   413
  val trans_tac         = trans_tac
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   414
  val norm_tac =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   415
     ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps @ mult_1s))
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   416
     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@real_mult_minus_simps))
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14321
diff changeset
   417
     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   418
  val numeral_simp_tac  =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   419
         ALLGOALS (simp_tac (HOL_ss addsimps rel_real_number_of@bin_simps))
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   420
  val simplify_meta_eq  = simplify_meta_eq
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   421
  end
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   422
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   423
structure DivCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   424
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   425
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   426
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   427
  val dest_bal = HOLogic.dest_bin "HOL.divide" HOLogic.realT
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14293
diff changeset
   428
  val cancel = mult_divide_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   429
  val neg_exchanges = false
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   430
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   431
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   432
structure EqCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   433
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   434
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   435
  val mk_bal   = HOLogic.mk_eq
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   436
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14293
diff changeset
   437
  val cancel = mult_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   438
  val neg_exchanges = false
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   439
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   440
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   441
structure LessCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   442
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   443
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   444
  val mk_bal   = HOLogic.mk_binrel "op <"
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   445
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14293
diff changeset
   446
  val cancel = mult_less_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   447
  val neg_exchanges = true
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   448
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   449
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   450
structure LeCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   451
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   452
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   453
  val mk_bal   = HOLogic.mk_binrel "op <="
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   454
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14293
diff changeset
   455
  val cancel = mult_le_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   456
  val neg_exchanges = true
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   457
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   458
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   459
val real_cancel_numeral_factors_relations =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   460
  map prep_simproc
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   461
   [("realeq_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   462
     ["(l::real) * m = n", "(l::real) = m * n"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   463
     EqCancelNumeralFactor.proc),
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   464
    ("realless_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   465
     ["(l::real) * m < n", "(l::real) < m * n"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   466
     LessCancelNumeralFactor.proc),
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   467
    ("realle_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   468
     ["(l::real) * m <= n", "(l::real) <= m * n"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   469
     LeCancelNumeralFactor.proc)]
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   470
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   471
val real_cancel_numeral_factors_divide = prep_simproc
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   472
        ("realdiv_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   473
         ["((l::real) * m) / n", "(l::real) / (m * n)",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   474
          "((number_of v)::real) / (number_of w)"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   475
         DivCancelNumeralFactor.proc)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   476
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   477
val real_cancel_numeral_factors =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   478
    real_cancel_numeral_factors_relations @
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   479
    [real_cancel_numeral_factors_divide]
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   480
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   481
end;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   482
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   483
Addsimprocs real_cancel_numeral_factors;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   484
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   485
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   486
(*examples:
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   487
print_depth 22;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   488
set timing;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   489
set trace_simp;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   490
fun test s = (Goal s; by (Simp_tac 1));
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   491
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   492
test "0 <= (y::real) * -2";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   493
test "9*x = 12 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   494
test "(9*x) / (12 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   495
test "9*x < 12 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   496
test "9*x <= 12 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   497
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   498
test "-99*x = 132 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   499
test "(-99*x) / (132 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   500
test "-99*x < 132 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   501
test "-99*x <= 132 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   502
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   503
test "999*x = -396 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   504
test "(999*x) / (-396 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   505
test "999*x < -396 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   506
test "999*x <= -396 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   507
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   508
test  "(- ((2::real) * x) <= 2 * y)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   509
test "-99*x = -81 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   510
test "(-99*x) / (-81 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   511
test "-99*x <= -81 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   512
test "-99*x < -81 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   513
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   514
test "-2 * x = -1 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   515
test "-2 * x = -(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   516
test "(-2 * x) / (-1 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   517
test "-2 * x < -(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   518
test "-2 * x <= -1 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   519
test "-x < -23 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   520
test "-x <= -23 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   521
*)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   522
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   523
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   524
(** Declarations for ExtractCommonTerm **)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   525
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   526
local
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   527
  open Real_Numeral_Simprocs
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   528
in
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   529
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   530
structure CancelFactorCommon =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   531
  struct
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   532
  val mk_sum            = long_mk_prod
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   533
  val dest_sum          = dest_prod
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   534
  val mk_coeff          = mk_coeff
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   535
  val dest_coeff        = dest_coeff
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   536
  val find_first        = find_first []
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   537
  val trans_tac         = trans_tac
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14321
diff changeset
   538
  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   539
  end;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   540
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   541
structure EqCancelFactor = ExtractCommonTermFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   542
 (open CancelFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   543
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   544
  val mk_bal   = HOLogic.mk_eq
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   545
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14293
diff changeset
   546
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_cancel_left
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   547
);
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   548
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   549
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   550
structure DivideCancelFactor = ExtractCommonTermFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   551
 (open CancelFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   552
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   553
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   554
  val dest_bal = HOLogic.dest_bin "HOL.divide" HOLogic.realT
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14293
diff changeset
   555
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   556
);
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   557
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   558
val real_cancel_factor =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   559
  map prep_simproc
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   560
   [("real_eq_cancel_factor", ["(l::real) * m = n", "(l::real) = m * n"], EqCancelFactor.proc),
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   561
    ("real_divide_cancel_factor", ["((l::real) * m) / n", "(l::real) / (m * n)"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   562
     DivideCancelFactor.proc)];
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   563
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   564
end;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   565
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   566
Addsimprocs real_cancel_factor;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   567
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   568
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   569
(*examples:
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   570
print_depth 22;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   571
set timing;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   572
set trace_simp;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   573
fun test s = (Goal s; by (Asm_simp_tac 1));
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   574
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   575
test "x*k = k*(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   576
test "k = k*(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   577
test "a*(b*c) = (b::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   578
test "a*(b*c) = d*(b::real)*(x*a)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   579
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   580
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   581
test "(x*k) / (k*(y::real)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   582
test "(k) / (k*(y::real)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   583
test "(a*(b*c)) / ((b::real)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   584
test "(a*(b*c)) / (d*(b::real)*(x*a)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   585
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   586
(*FIXME: what do we do about this?*)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   587
test "a*(b*c)/(y*z) = d*(b::real)*(x*a)/z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   588
*)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   589
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   590
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   591
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   592
(****Instantiation of the generic linear arithmetic package****)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   593
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   594
local
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   595
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   596
(* reduce contradictory <= to False *)
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   597
val add_rules = 
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14368
diff changeset
   598
    [real_numeral_0_eq_0, real_numeral_1_eq_1,
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   599
     add_real_number_of, minus_real_number_of, diff_real_number_of,
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14368
diff changeset
   600
     mult_real_number_of, eq_real_number_of, less_real_number_of];
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   601
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   602
val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   603
                real_cancel_numeral_factors_divide]@
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   604
               Real_Numeral_Simprocs.cancel_numerals @
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   605
               Real_Numeral_Simprocs.eval_numerals;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   606
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   607
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   608
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   609
val real_mult_mono_thms =
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   610
 [(rotate_prems 1 real_mult_less_mono2,
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   611
   cvar(real_mult_less_mono2, hd(prems_of real_mult_less_mono2))),
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14321
diff changeset
   612
  (real_mult_left_mono,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14321
diff changeset
   613
   cvar(real_mult_left_mono, hd(tl(prems_of real_mult_left_mono))))]
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   614
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14368
diff changeset
   615
val simps = [real_of_nat_zero, real_of_nat_Suc, real_of_nat_add, 
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14368
diff changeset
   616
       real_of_nat_mult, real_of_int_zero, real_of_one, real_of_int_add RS sym,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14368
diff changeset
   617
       real_of_int_minus RS sym, real_of_int_diff RS sym,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14368
diff changeset
   618
       real_of_int_mult RS sym, symmetric real_number_of_def];
14355
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   619
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   620
val int_inj_thms = [real_of_int_le_iff RS iffD2, real_of_int_less_iff RS iffD2,
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   621
                    real_of_int_inject RS iffD2];
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   622
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   623
val nat_inj_thms = [real_of_nat_le_iff RS iffD2, real_of_nat_less_iff RS iffD2,
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   624
                    real_of_nat_inject RS iffD2];
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   625
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   626
in
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   627
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   628
val fast_real_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   629
  "fast_real_arith" ["(m::real) < n","(m::real) <= n", "(m::real) = n"]
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   630
  Fast_Arith.lin_arith_prover;
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   631
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   632
val real_arith_setup =
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   633
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
14368
2763da611ad9 converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents: 14365
diff changeset
   634
   {add_mono_thms = add_mono_thms,
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   635
    mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
14355
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   636
    inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms,
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14356
diff changeset
   637
    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the reals are dense!*)
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   638
    simpset = simpset addsimps add_rules
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   639
                      addsimps simps
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   640
                      addsimprocs simprocs}),
14355
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   641
  arith_inj_const ("RealDef.real", HOLogic.natT --> HOLogic.realT),
67e2e96bfe36 Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents: 14352
diff changeset
   642
  arith_inj_const ("RealDef.real", HOLogic.intT --> HOLogic.realT),
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   643
  arith_discrete ("RealDef.real",false),
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   644
  Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]];
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   645
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   646
(* some thms for injection nat => real:
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   647
real_of_nat_zero
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   648
?zero_eq_numeral_0
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   649
real_of_nat_add
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   650
*)
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   651
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   652
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   653
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   654
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   655
(* Some test data [omitting examples that assume the ordering to be discrete!]
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   656
Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   657
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   658
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   659
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   660
Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   661
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   662
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   663
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   664
Goal "!!a::real. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   665
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   666
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   667
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   668
Goal "!!a::real. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   669
by (arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   670
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   671
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   672
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   673
\     ==> a <= l";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   674
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   675
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   676
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   677
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   678
\     ==> a+a+a+a <= l+l+l+l";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   679
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   680
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   681
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   682
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   683
\     ==> a+a+a+a+a <= l+l+l+l+i";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   684
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   685
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   686
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   687
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   688
\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   689
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   690
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   691
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   692
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   693
\     ==> 6*a <= 5*l+i";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   694
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   695
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   696
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   697
Goal "a<=b ==> a < b+(1::real)";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   698
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   699
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   700
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   701
Goal "a<=b ==> a-(3::real) < b";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   702
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   703
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   704
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   705
Goal "a<=b ==> a-(1::real) < b";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   706
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   707
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   708
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   709
*)
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   710
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   711
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   712