src/HOL/Hyperreal/hypreal_arith.ML
author paulson
Thu, 29 Jan 2004 16:51:17 +0100
changeset 14370 b0064703967b
parent 14369 c50188fe6366
child 14387 e96d5c42c4b0
permissions -rw-r--r--
simplifications in the hyperreals
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
     1
(*  Title:      HOL/Hyperreal/hypreal_arith.ML
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     2
    ID:         $Id$
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     3
    Author:     Tobias Nipkow, TU Muenchen
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   1999 TU Muenchen
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     5
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
     6
Simprocs for common factor cancellation & Rational coefficient handling
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
     7
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
     8
Instantiation of the generic linear arithmetic package for type hypreal.
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     9
*)
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    10
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    11
(*FIXME DELETE*)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
    12
val hypreal_mult_less_mono2 = thm"hypreal_mult_less_mono2";
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
    13
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    14
val hypreal_mult_left_mono =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    15
    read_instantiate_sg(sign_of (the_context())) [("a","?a::hypreal")] mult_left_mono;
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    16
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    17
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    18
val hypreal_number_of = thm "hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    19
val hypreal_numeral_0_eq_0 = thm "hypreal_numeral_0_eq_0";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    20
val hypreal_numeral_1_eq_1 = thm "hypreal_numeral_1_eq_1";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    21
val hypreal_number_of_def = thm "hypreal_number_of_def";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    22
val add_hypreal_number_of = thm "add_hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    23
val minus_hypreal_number_of = thm "minus_hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    24
val diff_hypreal_number_of = thm "diff_hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    25
val mult_hypreal_number_of = thm "mult_hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    26
val hypreal_mult_2 = thm "hypreal_mult_2";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    27
val hypreal_mult_2_right = thm "hypreal_mult_2_right";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    28
val eq_hypreal_number_of = thm "eq_hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    29
val less_hypreal_number_of = thm "less_hypreal_number_of";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    30
val hypreal_minus_1_eq_m1 = thm "hypreal_minus_1_eq_m1";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    31
val hypreal_mult_minus1 = thm "hypreal_mult_minus1";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    32
val hypreal_mult_minus1_right = thm "hypreal_mult_minus1_right";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    33
val hypreal_add_number_of_left = thm "hypreal_add_number_of_left";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    34
val hypreal_mult_number_of_left = thm "hypreal_mult_number_of_left";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    35
val hypreal_add_number_of_diff1 = thm "hypreal_add_number_of_diff1";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    36
val hypreal_add_number_of_diff2 = thm "hypreal_add_number_of_diff2";
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    37
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    38
(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    39
val hypreal_numeral_ss =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    40
    real_numeral_ss addsimps [hypreal_numeral_0_eq_0 RS sym,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    41
                              hypreal_numeral_1_eq_1 RS sym,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    42
                              hypreal_minus_1_eq_m1]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    43
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    44
fun rename_numerals th =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    45
    asm_full_simplify hypreal_numeral_ss (Thm.transfer (the_context ()) th)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    46
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    47
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    48
structure Hyperreal_Numeral_Simprocs =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    49
struct
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    50
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    51
(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    52
  isn't complicated by the abstract 0 and 1.*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    53
val numeral_syms = [hypreal_numeral_0_eq_0 RS sym,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    54
                    hypreal_numeral_1_eq_1 RS sym]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    55
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    56
(*Utilities*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    57
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    58
val hyprealT = Type("HyperDef.hypreal",[])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    59
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    60
fun mk_numeral n = HOLogic.number_of_const hyprealT $ HOLogic.mk_bin n
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    61
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    62
val dest_numeral = Real_Numeral_Simprocs.dest_numeral
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    63
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    64
val find_first_numeral = Real_Numeral_Simprocs.find_first_numeral
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    65
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    66
val zero = mk_numeral 0
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    67
val mk_plus = Real_Numeral_Simprocs.mk_plus
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    68
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    69
val uminus_const = Const ("uminus", hyprealT --> hyprealT)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    70
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    71
(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    72
fun mk_sum []        = zero
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    73
  | mk_sum [t,u]     = mk_plus (t, u)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    74
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    75
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    76
(*this version ALWAYS includes a trailing zero*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    77
fun long_mk_sum []        = zero
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    78
  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    79
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    80
val dest_plus = HOLogic.dest_bin "op +" hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    81
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    82
(*decompose additions AND subtractions as a sum*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    83
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    84
        dest_summing (pos, t, dest_summing (pos, u, ts))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    85
  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    86
        dest_summing (pos, t, dest_summing (not pos, u, ts))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    87
  | dest_summing (pos, t, ts) =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    88
        if pos then t::ts else uminus_const$t :: ts
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    89
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    90
fun dest_sum t = dest_summing (true, t, [])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    91
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    92
val mk_diff = HOLogic.mk_binop "op -"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    93
val dest_diff = HOLogic.dest_bin "op -" hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    94
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    95
val one = mk_numeral 1
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    96
val mk_times = HOLogic.mk_binop "op *"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    97
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    98
fun mk_prod [] = one
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    99
  | mk_prod [t] = t
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   100
  | mk_prod (t :: ts) = if t = one then mk_prod ts
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   101
                        else mk_times (t, mk_prod ts)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   102
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   103
val dest_times = HOLogic.dest_bin "op *" hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   104
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   105
fun dest_prod t =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   106
      let val (t,u) = dest_times t
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   107
      in  dest_prod t @ dest_prod u  end
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   108
      handle TERM _ => [t]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   109
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   110
(*DON'T do the obvious simplifications; that would create special cases*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   111
fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   112
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   113
(*Express t as a product of (possibly) a numeral with other sorted terms*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   114
fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   115
  | dest_coeff sign t =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   116
    let val ts = sort Term.term_ord (dest_prod t)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   117
        val (n, ts') = find_first_numeral [] ts
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   118
                          handle TERM _ => (1, ts)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   119
    in (sign*n, mk_prod ts') end
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   120
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   121
(*Find first coefficient-term THAT MATCHES u*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   122
fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   123
  | find_first_coeff past u (t::terms) =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   124
        let val (n,u') = dest_coeff 1 t
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   125
        in  if u aconv u' then (n, rev past @ terms)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   126
                          else find_first_coeff (t::past) u terms
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   127
        end
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   128
        handle TERM _ => find_first_coeff (t::past) u terms
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   129
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   130
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   131
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   132
val add_0s = map rename_numerals
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   133
                 [hypreal_add_zero_left, hypreal_add_zero_right]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   134
val mult_1s = map rename_numerals [hypreal_mult_1, hypreal_mult_1_right] @
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   135
              [hypreal_mult_minus1, hypreal_mult_minus1_right]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   136
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   137
(*To perform binary arithmetic*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   138
val bin_simps =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   139
    [hypreal_numeral_0_eq_0 RS sym, hypreal_numeral_1_eq_1 RS sym,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   140
     add_hypreal_number_of, hypreal_add_number_of_left,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   141
     minus_hypreal_number_of,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   142
     diff_hypreal_number_of, mult_hypreal_number_of,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   143
     hypreal_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   144
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   145
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   146
  during re-arrangement*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   147
val non_add_bin_simps = 
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   148
    bin_simps \\ [hypreal_add_number_of_left, add_hypreal_number_of]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   149
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   150
(*To evaluate binary negations of coefficients*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   151
val hypreal_minus_simps = NCons_simps @
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   152
                   [hypreal_minus_1_eq_m1, minus_hypreal_number_of,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   153
                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   154
                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   155
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   156
(*To let us treat subtraction as addition*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   157
val diff_simps = [hypreal_diff_def, minus_add_distrib, minus_minus]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   158
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   159
(*push the unary minus down: - x * y = x * - y *)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   160
val hypreal_minus_mult_eq_1_to_2 =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   161
    [minus_mult_left RS sym, minus_mult_right] MRS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   162
    |> standard
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   163
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   164
(*to extract again any uncancelled minuses*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   165
val hypreal_minus_from_mult_simps =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   166
    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   167
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   168
(*combine unary minus with numeric literals, however nested within a product*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   169
val hypreal_mult_minus_simps =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   170
    [hypreal_mult_assoc, minus_mult_left, hypreal_minus_mult_eq_1_to_2]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   171
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   172
(*Final simplification: cancel + and *  *)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   173
val simplify_meta_eq =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   174
    Int_Numeral_Simprocs.simplify_meta_eq
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   175
         [hypreal_add_zero_left, hypreal_add_zero_right,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   176
          mult_zero_left, mult_zero_right, mult_1, mult_1_right]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   177
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   178
val prep_simproc = Real_Numeral_Simprocs.prep_simproc
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   179
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   180
structure CancelNumeralsCommon =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   181
  struct
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   182
  val mk_sum            = mk_sum
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   183
  val dest_sum          = dest_sum
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   184
  val mk_coeff          = mk_coeff
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   185
  val dest_coeff        = dest_coeff 1
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   186
  val find_first_coeff  = find_first_coeff []
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   187
  val trans_tac         = Real_Numeral_Simprocs.trans_tac
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   188
  val norm_tac =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   189
     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   190
                                         hypreal_minus_simps@add_ac))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   191
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hypreal_mult_minus_simps))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   192
     THEN ALLGOALS
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   193
              (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps@
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   194
                                         add_ac@mult_ac))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   195
  val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   196
  val simplify_meta_eq  = simplify_meta_eq
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   197
  end
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   198
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   199
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   200
structure EqCancelNumerals = CancelNumeralsFun
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   201
 (open CancelNumeralsCommon
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   202
  val prove_conv = Bin_Simprocs.prove_conv
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   203
  val mk_bal   = HOLogic.mk_eq
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   204
  val dest_bal = HOLogic.dest_bin "op =" hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   205
  val bal_add1 = eq_add_iff1 RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   206
  val bal_add2 = eq_add_iff2 RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   207
)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   208
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   209
structure LessCancelNumerals = CancelNumeralsFun
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   210
 (open CancelNumeralsCommon
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   211
  val prove_conv = Bin_Simprocs.prove_conv
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   212
  val mk_bal   = HOLogic.mk_binrel "op <"
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   213
  val dest_bal = HOLogic.dest_bin "op <" hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   214
  val bal_add1 = less_add_iff1 RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   215
  val bal_add2 = less_add_iff2 RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   216
)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   217
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   218
structure LeCancelNumerals = CancelNumeralsFun
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   219
 (open CancelNumeralsCommon
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   220
  val prove_conv = Bin_Simprocs.prove_conv
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   221
  val mk_bal   = HOLogic.mk_binrel "op <="
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   222
  val dest_bal = HOLogic.dest_bin "op <=" hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   223
  val bal_add1 = le_add_iff1 RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   224
  val bal_add2 = le_add_iff2 RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   225
)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   226
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   227
val cancel_numerals =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   228
  map prep_simproc
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   229
   [("hyprealeq_cancel_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   230
     ["(l::hypreal) + m = n", "(l::hypreal) = m + n",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   231
      "(l::hypreal) - m = n", "(l::hypreal) = m - n",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   232
      "(l::hypreal) * m = n", "(l::hypreal) = m * n"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   233
     EqCancelNumerals.proc),
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   234
    ("hyprealless_cancel_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   235
     ["(l::hypreal) + m < n", "(l::hypreal) < m + n",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   236
      "(l::hypreal) - m < n", "(l::hypreal) < m - n",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   237
      "(l::hypreal) * m < n", "(l::hypreal) < m * n"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   238
     LessCancelNumerals.proc),
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   239
    ("hyprealle_cancel_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   240
     ["(l::hypreal) + m <= n", "(l::hypreal) <= m + n",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   241
      "(l::hypreal) - m <= n", "(l::hypreal) <= m - n",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   242
      "(l::hypreal) * m <= n", "(l::hypreal) <= m * n"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   243
     LeCancelNumerals.proc)]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   244
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   245
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   246
structure CombineNumeralsData =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   247
  struct
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   248
  val add               = op + : int*int -> int
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   249
  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   250
  val dest_sum          = dest_sum
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   251
  val mk_coeff          = mk_coeff
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   252
  val dest_coeff        = dest_coeff 1
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   253
  val left_distrib      = combine_common_factor RS trans
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   254
  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   255
  val trans_tac         = Real_Numeral_Simprocs.trans_tac
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   256
  val norm_tac =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   257
     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   258
                                   diff_simps@hypreal_minus_simps@add_ac))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   259
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hypreal_mult_minus_simps))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   260
     THEN ALLGOALS (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps@
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   261
                                              add_ac@mult_ac))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   262
  val numeral_simp_tac  = ALLGOALS
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   263
                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   264
  val simplify_meta_eq  = simplify_meta_eq
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   265
  end
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   266
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   267
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   268
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   269
val combine_numerals =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   270
  prep_simproc
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   271
    ("hypreal_combine_numerals", ["(i::hypreal) + j", "(i::hypreal) - j"], CombineNumerals.proc)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   272
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   273
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   274
(** Declarations for ExtractCommonTerm **)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   275
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   276
(*this version ALWAYS includes a trailing one*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   277
fun long_mk_prod []        = one
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   278
  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   279
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   280
(*Find first term that matches u*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   281
fun find_first past u []         = raise TERM("find_first", [])
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   282
  | find_first past u (t::terms) =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   283
        if u aconv t then (rev past @ terms)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   284
        else find_first (t::past) u terms
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   285
        handle TERM _ => find_first (t::past) u terms
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   286
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   287
(*Final simplification: cancel + and *  *)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   288
fun cancel_simplify_meta_eq cancel_th th =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   289
    Int_Numeral_Simprocs.simplify_meta_eq
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   290
        [mult_1, mult_1_right]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   291
        (([th, cancel_th]) MRS trans)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   292
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   293
(*** Making constant folding work for 0 and 1 too ***)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   294
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   295
structure HyperrealAbstractNumeralsData =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   296
  struct
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   297
  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   298
  val is_numeral      = Bin_Simprocs.is_numeral
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   299
  val numeral_0_eq_0  = hypreal_numeral_0_eq_0
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   300
  val numeral_1_eq_1  = hypreal_numeral_1_eq_1
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   301
  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   302
  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   303
  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   304
  end
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   305
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   306
structure HyperrealAbstractNumerals =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   307
  AbstractNumeralsFun (HyperrealAbstractNumeralsData)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   308
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   309
(*For addition, we already have rules for the operand 0.
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   310
  Multiplication is omitted because there are already special rules for
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   311
  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   312
  For the others, having three patterns is a compromise between just having
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   313
  one (many spurious calls) and having nine (just too many!) *)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   314
val eval_numerals =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   315
  map prep_simproc
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   316
   [("hypreal_add_eval_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   317
     ["(m::hypreal) + 1", "(m::hypreal) + number_of v"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   318
     HyperrealAbstractNumerals.proc add_hypreal_number_of),
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   319
    ("hypreal_diff_eval_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   320
     ["(m::hypreal) - 1", "(m::hypreal) - number_of v"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   321
     HyperrealAbstractNumerals.proc diff_hypreal_number_of),
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   322
    ("hypreal_eq_eval_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   323
     ["(m::hypreal) = 0", "(m::hypreal) = 1", "(m::hypreal) = number_of v"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   324
     HyperrealAbstractNumerals.proc eq_hypreal_number_of),
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   325
    ("hypreal_less_eval_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   326
     ["(m::hypreal) < 0", "(m::hypreal) < 1", "(m::hypreal) < number_of v"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   327
     HyperrealAbstractNumerals.proc less_hypreal_number_of),
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   328
    ("hypreal_le_eval_numerals",
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   329
     ["(m::hypreal) <= 0", "(m::hypreal) <= 1", "(m::hypreal) <= number_of v"],
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   330
     HyperrealAbstractNumerals.proc le_number_of_eq_not_less)]
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   331
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   332
end;
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   333
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   334
Addsimprocs Hyperreal_Numeral_Simprocs.eval_numerals;
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   335
Addsimprocs Hyperreal_Numeral_Simprocs.cancel_numerals;
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   336
Addsimprocs [Hyperreal_Numeral_Simprocs.combine_numerals];
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   337
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   338
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   339
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   340
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   341
(**** Constant folding for hypreal plus and times ****)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   342
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   343
(*We do not need
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   344
    structure Hyperreal_Plus_Assoc = Assoc_Fold (Hyperreal_Plus_Assoc_Data)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   345
  because combine_numerals does the same thing*)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   346
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   347
structure Hyperreal_Times_Assoc_Data : ASSOC_FOLD_DATA =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   348
struct
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   349
  val ss                = HOL_ss
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   350
  val eq_reflection     = eq_reflection
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   351
  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   352
  val T      = Hyperreal_Numeral_Simprocs.hyprealT
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   353
  val plus   = Const ("op *", [T,T] ---> T)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   354
  val add_ac = mult_ac
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   355
end;
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   356
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   357
structure Hyperreal_Times_Assoc = Assoc_Fold (Hyperreal_Times_Assoc_Data);
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   358
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   359
Addsimprocs [Hyperreal_Times_Assoc.conv];
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   360
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   361
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   362
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   363
(**** Simprocs for numeric literals ****)
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   364
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   365
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   366
(****Common factor cancellation****)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   367
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   368
local
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   369
  open Hyperreal_Numeral_Simprocs
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   370
in
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   371
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   372
val rel_hypreal_number_of = [eq_hypreal_number_of, less_hypreal_number_of,
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   373
                             le_number_of_eq_not_less];
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   374
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   375
structure CancelNumeralFactorCommon =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   376
  struct
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   377
  val mk_coeff          = mk_coeff
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   378
  val dest_coeff        = dest_coeff 1
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   379
  val trans_tac         = Real_Numeral_Simprocs.trans_tac
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   380
  val norm_tac =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   381
     ALLGOALS (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps @ mult_1s))
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   382
     THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps))
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14309
diff changeset
   383
     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   384
  val numeral_simp_tac  =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   385
         ALLGOALS (simp_tac (HOL_ss addsimps rel_hypreal_number_of@bin_simps))
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   386
  val simplify_meta_eq  = simplify_meta_eq
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   387
  end
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   388
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   389
structure DivCancelNumeralFactor = CancelNumeralFactorFun
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   390
 (open CancelNumeralFactorCommon
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   391
  val prove_conv = Bin_Simprocs.prove_conv
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   392
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   393
  val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   394
  val cancel = mult_divide_cancel_left RS trans
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   395
  val neg_exchanges = false
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   396
)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   397
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   398
structure EqCancelNumeralFactor = CancelNumeralFactorFun
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   399
 (open CancelNumeralFactorCommon
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   400
  val prove_conv = Bin_Simprocs.prove_conv
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   401
  val mk_bal   = HOLogic.mk_eq
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   402
  val dest_bal = HOLogic.dest_bin "op =" hyprealT
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   403
  val cancel = mult_cancel_left RS trans
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   404
  val neg_exchanges = false
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   405
)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   406
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   407
structure LessCancelNumeralFactor = CancelNumeralFactorFun
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   408
 (open CancelNumeralFactorCommon
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   409
  val prove_conv = Bin_Simprocs.prove_conv
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   410
  val mk_bal   = HOLogic.mk_binrel "op <"
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   411
  val dest_bal = HOLogic.dest_bin "op <" hyprealT
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   412
  val cancel = mult_less_cancel_left RS trans
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   413
  val neg_exchanges = true
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   414
)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   415
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   416
structure LeCancelNumeralFactor = CancelNumeralFactorFun
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   417
 (open CancelNumeralFactorCommon
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   418
  val prove_conv = Bin_Simprocs.prove_conv
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   419
  val mk_bal   = HOLogic.mk_binrel "op <="
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   420
  val dest_bal = HOLogic.dest_bin "op <=" hyprealT
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   421
  val cancel = mult_le_cancel_left RS trans
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   422
  val neg_exchanges = true
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   423
)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   424
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   425
val hypreal_cancel_numeral_factors_relations =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   426
  map prep_simproc
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   427
   [("hyprealeq_cancel_numeral_factor",
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   428
     ["(l::hypreal) * m = n", "(l::hypreal) = m * n"],
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   429
     EqCancelNumeralFactor.proc),
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   430
    ("hyprealless_cancel_numeral_factor",
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   431
     ["(l::hypreal) * m < n", "(l::hypreal) < m * n"],
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   432
     LessCancelNumeralFactor.proc),
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   433
    ("hyprealle_cancel_numeral_factor",
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   434
     ["(l::hypreal) * m <= n", "(l::hypreal) <= m * n"],
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   435
     LeCancelNumeralFactor.proc)];
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   436
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   437
val hypreal_cancel_numeral_factors_divide = prep_simproc
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   438
        ("hyprealdiv_cancel_numeral_factor",
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   439
         ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)",
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   440
          "((number_of v)::hypreal) / (number_of w)"],
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   441
         DivCancelNumeralFactor.proc);
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   442
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   443
val hypreal_cancel_numeral_factors =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   444
    hypreal_cancel_numeral_factors_relations @
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   445
    [hypreal_cancel_numeral_factors_divide];
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   446
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   447
end;
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   448
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   449
Addsimprocs hypreal_cancel_numeral_factors;
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   450
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   451
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   452
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   453
(** Declarations for ExtractCommonTerm **)
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   454
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   455
local
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   456
  open Hyperreal_Numeral_Simprocs
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   457
in
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   458
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   459
structure CancelFactorCommon =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   460
  struct
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   461
  val mk_sum            = long_mk_prod
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   462
  val dest_sum          = dest_prod
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   463
  val mk_coeff          = mk_coeff
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   464
  val dest_coeff        = dest_coeff
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   465
  val find_first        = find_first []
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   466
  val trans_tac         = Real_Numeral_Simprocs.trans_tac
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14309
diff changeset
   467
  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac))
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   468
  end;
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   469
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   470
structure EqCancelFactor = ExtractCommonTermFun
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   471
 (open CancelFactorCommon
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   472
  val prove_conv = Bin_Simprocs.prove_conv
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   473
  val mk_bal   = HOLogic.mk_eq
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   474
  val dest_bal = HOLogic.dest_bin "op =" hyprealT
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   475
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_cancel_left
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   476
);
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   477
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   478
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   479
structure DivideCancelFactor = ExtractCommonTermFun
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   480
 (open CancelFactorCommon
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   481
  val prove_conv = Bin_Simprocs.prove_conv
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   482
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   483
  val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   484
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_divide_cancel_eq_if
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   485
);
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   486
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   487
val hypreal_cancel_factor =
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   488
  map prep_simproc
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   489
   [("hypreal_eq_cancel_factor", ["(l::hypreal) * m = n", "(l::hypreal) = m * n"],
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   490
     EqCancelFactor.proc),
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   491
    ("hypreal_divide_cancel_factor", ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)"],
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   492
     DivideCancelFactor.proc)];
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   493
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   494
end;
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   495
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   496
Addsimprocs hypreal_cancel_factor;
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   497
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   498
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   499
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   500
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   501
(****Instantiation of the generic linear arithmetic package****)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   502
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
   503
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   504
local
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   505
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   506
(* reduce contradictory <= to False *)
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   507
val add_rules =
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   508
    [hypreal_numeral_0_eq_0, hypreal_numeral_1_eq_1,
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   509
     add_hypreal_number_of, minus_hypreal_number_of, diff_hypreal_number_of,
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   510
     mult_hypreal_number_of, eq_hypreal_number_of, less_hypreal_number_of];
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   511
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   512
val simprocs = [Hyperreal_Times_Assoc.conv, 
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   513
                Hyperreal_Numeral_Simprocs.combine_numerals,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   514
                hypreal_cancel_numeral_factors_divide]@
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   515
               Hyperreal_Numeral_Simprocs.cancel_numerals @
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   516
               Hyperreal_Numeral_Simprocs.eval_numerals;
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   517
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   518
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   519
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   520
val hypreal_mult_mono_thms =
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   521
 [(rotate_prems 1 hypreal_mult_less_mono2,
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   522
   cvar(hypreal_mult_less_mono2, hd(prems_of hypreal_mult_less_mono2))),
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   523
  (hypreal_mult_left_mono,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   524
   cvar(hypreal_mult_left_mono, hd(tl(prems_of hypreal_mult_left_mono))))]
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   525
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   526
val real_inj_thms = [hypreal_of_real_le_iff RS iffD2, 
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   527
                     hypreal_of_real_less_iff RS iffD2,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   528
                     hypreal_of_real_eq_iff RS iffD2];
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   529
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   530
in
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   531
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   532
val fast_hypreal_arith_simproc =
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   533
    Simplifier.simproc (Theory.sign_of (the_context ()))
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   534
      "fast_hypreal_arith" 
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   535
      ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"]
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   536
    Fast_Arith.lin_arith_prover;
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   537
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   538
val hypreal_arith_setup =
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   539
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   540
   {add_mono_thms = add_mono_thms,
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   541
    mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   542
    inj_thms = inj_thms @ real_inj_thms, 
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   543
    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the hypreals are dense!*)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   544
    simpset = simpset addsimps add_rules
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   545
                      addsimprocs simprocs}),
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   546
  arith_inj_const ("HyperDef.hypreal_of_real", 
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
   547
                   HOLogic.realT --> Hyperreal_Numeral_Simprocs.hyprealT),
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   548
  arith_discrete ("HyperDef.hypreal",false),
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   549
  Simplifier.change_simpset_of (op addsimprocs) [fast_hypreal_arith_simproc]];
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   550
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   551
end;
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   552
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   553
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
   554