src/HOL/Real/real_arith.ML
author paulson
Mon, 12 Jan 2004 16:51:45 +0100
changeset 14353 79f9fbef9106
parent 14352 a8b1a44d8264
child 14355 67e2e96bfe36
permissions -rw-r--r--
Added lemmas to Ring_and_Field with slightly modified simplification rules Deleted some little-used integer theorems, replacing them by the generic ones in Ring_and_Field Consolidated integer powers
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
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    11
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    12
(****Common factor cancellation****)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    13
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    14
(*To quote from Provers/Arith/cancel_numeral_factor.ML:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    15
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    16
This simproc Cancels common coefficients in balanced expressions:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    17
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    18
     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    19
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    20
where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    21
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
    22
*)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14289
diff changeset
    23
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    24
local
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    25
  open Real_Numeral_Simprocs
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    26
in
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    27
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    28
val rel_real_number_of = [eq_real_number_of, less_real_number_of,
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    29
                          le_real_number_of_eq_not_less]
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    30
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    31
structure CancelNumeralFactorCommon =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    32
  struct
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    33
  val mk_coeff          = mk_coeff
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    34
  val dest_coeff        = dest_coeff 1
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    35
  val trans_tac         = trans_tac
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    36
  val norm_tac =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    37
     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
    38
     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
    39
     THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac))
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    40
  val numeral_simp_tac  =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    41
         ALLGOALS (simp_tac (HOL_ss addsimps rel_real_number_of@bin_simps))
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    42
  val simplify_meta_eq  = simplify_meta_eq
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    43
  end
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    44
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    45
structure DivCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    46
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    47
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    48
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    49
  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
    50
  val cancel = mult_divide_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    51
  val neg_exchanges = false
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    52
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    53
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    54
structure EqCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    55
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    56
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    57
  val mk_bal   = HOLogic.mk_eq
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    58
  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
    59
  val cancel = mult_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    60
  val neg_exchanges = false
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    61
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    62
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    63
structure LessCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    64
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    65
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    66
  val mk_bal   = HOLogic.mk_binrel "op <"
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    67
  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
    68
  val cancel = mult_less_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    69
  val neg_exchanges = true
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    70
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    71
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    72
structure LeCancelNumeralFactor = CancelNumeralFactorFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    73
 (open CancelNumeralFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    74
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    75
  val mk_bal   = HOLogic.mk_binrel "op <="
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    76
  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
    77
  val cancel = mult_le_cancel_left RS trans
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    78
  val neg_exchanges = true
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    79
)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    80
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    81
val real_cancel_numeral_factors_relations =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    82
  map prep_simproc
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    83
   [("realeq_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    84
     ["(l::real) * m = n", "(l::real) = m * n"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    85
     EqCancelNumeralFactor.proc),
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    86
    ("realless_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    87
     ["(l::real) * m < n", "(l::real) < m * n"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    88
     LessCancelNumeralFactor.proc),
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    89
    ("realle_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    90
     ["(l::real) * m <= n", "(l::real) <= m * n"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    91
     LeCancelNumeralFactor.proc)]
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    92
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    93
val real_cancel_numeral_factors_divide = prep_simproc
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    94
        ("realdiv_cancel_numeral_factor",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    95
         ["((l::real) * m) / n", "(l::real) / (m * n)",
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    96
          "((number_of v)::real) / (number_of w)"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    97
         DivCancelNumeralFactor.proc)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    98
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
    99
val real_cancel_numeral_factors =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   100
    real_cancel_numeral_factors_relations @
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   101
    [real_cancel_numeral_factors_divide]
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   102
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   103
end;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   104
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   105
Addsimprocs real_cancel_numeral_factors;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   106
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   107
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   108
(*examples:
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   109
print_depth 22;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   110
set timing;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   111
set trace_simp;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   112
fun test s = (Goal s; by (Simp_tac 1));
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   113
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   114
test "0 <= (y::real) * -2";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   115
test "9*x = 12 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   116
test "(9*x) / (12 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   117
test "9*x < 12 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   118
test "9*x <= 12 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   119
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   120
test "-99*x = 132 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   121
test "(-99*x) / (132 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   122
test "-99*x < 132 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   123
test "-99*x <= 132 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   124
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   125
test "999*x = -396 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   126
test "(999*x) / (-396 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   127
test "999*x < -396 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   128
test "999*x <= -396 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   129
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   130
test  "(- ((2::real) * x) <= 2 * y)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   131
test "-99*x = -81 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   132
test "(-99*x) / (-81 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   133
test "-99*x <= -81 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   134
test "-99*x < -81 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   135
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   136
test "-2 * x = -1 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   137
test "-2 * x = -(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   138
test "(-2 * x) / (-1 * (y::real)) = z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   139
test "-2 * x < -(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   140
test "-2 * x <= -1 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   141
test "-x < -23 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   142
test "-x <= -23 * (y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   143
*)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   144
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   145
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   146
(** Declarations for ExtractCommonTerm **)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   147
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   148
local
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   149
  open Real_Numeral_Simprocs
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   150
in
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   151
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   152
structure CancelFactorCommon =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   153
  struct
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   154
  val mk_sum            = long_mk_prod
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   155
  val dest_sum          = dest_prod
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   156
  val mk_coeff          = mk_coeff
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   157
  val dest_coeff        = dest_coeff
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   158
  val find_first        = find_first []
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   159
  val trans_tac         = trans_tac
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14321
diff changeset
   160
  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
   161
  end;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   162
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   163
structure EqCancelFactor = ExtractCommonTermFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   164
 (open CancelFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   165
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   166
  val mk_bal   = HOLogic.mk_eq
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   167
  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
   168
  val simplify_meta_eq  = cancel_simplify_meta_eq mult_cancel_left
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   169
);
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   170
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   171
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   172
structure DivideCancelFactor = ExtractCommonTermFun
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   173
 (open CancelFactorCommon
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   174
  val prove_conv = Bin_Simprocs.prove_conv
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   175
  val mk_bal   = HOLogic.mk_binop "HOL.divide"
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   176
  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
   177
  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
   178
);
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   179
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   180
val real_cancel_factor =
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   181
  map prep_simproc
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   182
   [("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
   183
    ("real_divide_cancel_factor", ["((l::real) * m) / n", "(l::real) / (m * n)"],
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   184
     DivideCancelFactor.proc)];
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   185
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   186
end;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   187
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   188
Addsimprocs real_cancel_factor;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   189
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   190
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   191
(*examples:
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   192
print_depth 22;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   193
set timing;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   194
set trace_simp;
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   195
fun test s = (Goal s; by (Asm_simp_tac 1));
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   196
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   197
test "x*k = k*(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   198
test "k = k*(y::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   199
test "a*(b*c) = (b::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   200
test "a*(b*c) = d*(b::real)*(x*a)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   201
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   202
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   203
test "(x*k) / (k*(y::real)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   204
test "(k) / (k*(y::real)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   205
test "(a*(b*c)) / ((b::real)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   206
test "(a*(b*c)) / (d*(b::real)*(x*a)) = (uu::real)";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   207
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   208
(*FIXME: what do we do about this?*)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   209
test "a*(b*c)/(y*z) = d*(b::real)*(x*a)/z";
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   210
*)
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   211
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   212
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   213
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   214
(****Instantiation of the generic linear arithmetic package****)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   215
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   216
val add_zero_left = thm"Ring_and_Field.add_0";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   217
val add_zero_right = thm"Ring_and_Field.add_0_right";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   218
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   219
val real_mult_left_mono =
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   220
    read_instantiate_sg(sign_of RealBin.thy) [("a","?a::real")] mult_left_mono;
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   221
14289
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   222
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   223
local
deb8e1e62002 combining Real/{RealArith0,real_arith}.ML
paulson
parents: 14288
diff changeset
   224
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   225
(* reduce contradictory <= to False *)
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   226
val add_rules = 
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   227
    [order_less_irrefl, real_numeral_0_eq_0, real_numeral_1_eq_1,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   228
     real_minus_1_eq_m1, 
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   229
     add_real_number_of, minus_real_number_of, diff_real_number_of,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   230
     mult_real_number_of, eq_real_number_of, less_real_number_of,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   231
     le_real_number_of_eq_not_less, real_diff_def,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   232
     minus_add_distrib, minus_minus, mult_assoc, minus_zero,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   233
     add_zero_left, add_zero_right, left_minus, right_minus,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   234
     mult_zero_left, mult_zero_right, mult_1, mult_1_right,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   235
     minus_mult_left RS sym, minus_mult_right RS sym];
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   236
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   237
val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   238
                real_cancel_numeral_factors_divide]@
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   239
               Real_Numeral_Simprocs.cancel_numerals @
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   240
               Real_Numeral_Simprocs.eval_numerals;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   241
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   242
val mono_ss = simpset() addsimps
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   243
                [add_mono,add_strict_mono,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   244
                 real_add_less_le_mono,real_add_le_less_mono];
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   245
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   246
val add_mono_thms_real =
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   247
  map (fn s => prove_goal (the_context ()) s
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   248
                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   249
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   250
     "(i  = j) & (k <= l) ==> i + k <= j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   251
     "(i <= j) & (k  = l) ==> i + k <= j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   252
     "(i  = j) & (k  = l) ==> i + k  = j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   253
     "(i < j) & (k = l)   ==> i + k < j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   254
     "(i = j) & (k < l)   ==> i + k < j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   255
     "(i < j) & (k <= l)  ==> i + k < j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   256
     "(i <= j) & (k < l)  ==> i + k < j + (l::real)",
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   257
     "(i < j) & (k < l)   ==> i + k < j + (l::real)"];
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   258
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   259
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   260
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   261
val real_mult_mono_thms =
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   262
 [(rotate_prems 1 real_mult_less_mono2,
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   263
   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
   264
  (real_mult_left_mono,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14321
diff changeset
   265
   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
   266
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   267
(* reduce contradictory <= to False *)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   268
val simps = [True_implies_equals,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   269
             inst "a" "(number_of ?v)::real" right_distrib,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   270
             divide_1,times_divide_eq_right,times_divide_eq_left];
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   271
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   272
in
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   273
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   274
val fast_real_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   275
  "fast_real_arith" ["(m::real) < n","(m::real) <= n", "(m::real) = n"]
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   276
  Fast_Arith.lin_arith_prover;
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   277
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   278
val real_arith_setup =
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   279
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   280
   {add_mono_thms = add_mono_thms @ add_mono_thms_real,
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10574
diff changeset
   281
    mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   282
    inj_thms = inj_thms, (*FIXME: add real*)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   283
    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the reals are dense!*)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   284
    simpset = simpset addsimps add_rules
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   285
                      addsimps simps
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   286
                      addsimprocs simprocs}),
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   287
  arith_discrete ("RealDef.real",false),
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   288
  Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]];
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   289
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   290
(* some thms for injection nat => real:
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   291
real_of_nat_zero
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   292
?zero_eq_numeral_0
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   293
real_of_nat_add
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   294
*)
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   295
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   296
end;
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   297
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   298
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   299
(* Some test data [omitting examples that assume the ordering to be discrete!]
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   300
Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   301
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   302
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   303
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   304
Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   305
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   306
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   307
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   308
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
   309
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   310
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   311
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   312
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
   313
by (arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   314
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   315
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   316
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
   317
\     ==> a <= l";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   318
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   319
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   320
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   321
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
   322
\     ==> a+a+a+a <= l+l+l+l";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   323
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   324
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   325
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   326
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
   327
\     ==> a+a+a+a+a <= l+l+l+l+i";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   328
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   329
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   330
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   331
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
   332
\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   333
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   334
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   335
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   336
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
   337
\     ==> 6*a <= 5*l+i";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   338
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   339
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   340
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   341
Goal "a<=b ==> a < b+(1::real)";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   342
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   343
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   344
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   345
Goal "a<=b ==> a-(3::real) < b";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   346
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   347
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   348
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   349
Goal "a<=b ==> a-(1::real) < b";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   350
by (fast_arith_tac 1);
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   351
qed "";
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   352
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   353
*)
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   354
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   355
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14336
diff changeset
   356