src/HOL/Real/real_arith0.ML
author kleing
Tue, 13 May 2003 08:59:21 +0200
changeset 14024 213dcc39358f
parent 13554 4679359bb218
child 14321 55c688d2eefa
permissions -rw-r--r--
HOL-Real -> HOL-Complex
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Real/real_arith.ML
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     2
    ID:         $Id$
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow, TU Muenchen
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     4
    Copyright   1999 TU Muenchen
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     5
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     6
Instantiation of the generic linear arithmetic package for type real.
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     7
*)
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     8
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
     9
local
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    10
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    11
(* reduce contradictory <= to False *)
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    12
val add_rules = 
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    13
    [order_less_irrefl, real_numeral_0_eq_0, real_numeral_1_eq_1,
13554
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
    14
     real_minus_1_eq_m1, 
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    15
     add_real_number_of, minus_real_number_of, diff_real_number_of,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    16
     mult_real_number_of, eq_real_number_of, less_real_number_of,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    17
     le_real_number_of_eq_not_less, real_diff_def,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    18
     real_minus_add_distrib, real_minus_minus, real_mult_assoc,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    19
     real_minus_zero,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    20
     real_add_zero_left, real_add_zero_right,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    21
     real_add_minus, real_add_minus_left,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    22
     real_mult_0, real_mult_0_right,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    23
     real_mult_1, real_mult_1_right,
12483
0a01efff43e9 new rewrite rules for use by arith_tac to take care of uminus.
nipkow
parents: 12018
diff changeset
    24
     real_mult_minus_eq1, real_mult_minus_eq2];
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    25
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    26
val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals]@
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    27
               Real_Numeral_Simprocs.cancel_numerals @
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    28
               Real_Numeral_Simprocs.eval_numerals;
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    29
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    30
val mono_ss = simpset() addsimps
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    31
                [real_add_le_mono,real_add_less_mono,
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    32
                 real_add_less_le_mono,real_add_le_less_mono];
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    33
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    34
val add_mono_thms_real =
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    35
  map (fn s => prove_goal (the_context ()) s
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    36
                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    37
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    38
     "(i  = j) & (k <= l) ==> i + k <= j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    39
     "(i <= j) & (k  = l) ==> i + k <= j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    40
     "(i  = j) & (k  = l) ==> i + k  = j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    41
     "(i < j) & (k = l)   ==> i + k < j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    42
     "(i = j) & (k < l)   ==> i + k < j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    43
     "(i < j) & (k <= l)  ==> i + k < j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    44
     "(i <= j) & (k < l)  ==> i + k < j + (l::real)",
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    45
     "(i < j) & (k < l)   ==> i + k < j + (l::real)"];
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    46
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    47
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    48
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    49
val real_mult_mono_thms =
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    50
 [(rotate_prems 1 real_mult_less_mono2,
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    51
   cvar(real_mult_less_mono2, hd(prems_of real_mult_less_mono2))),
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    52
  (real_mult_le_mono2,
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    53
   cvar(real_mult_le_mono2, hd(tl(prems_of real_mult_le_mono2))))]
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    54
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    55
in
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    56
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12483
diff changeset
    57
val fast_real_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12483
diff changeset
    58
  "fast_real_arith" ["(m::real) < n","(m::real) <= n", "(m::real) = n"]
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12483
diff changeset
    59
  Fast_Arith.lin_arith_prover;
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    60
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    61
val real_arith_setup =
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    62
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    63
   {add_mono_thms = add_mono_thms @ add_mono_thms_real,
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    64
    mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    65
    inj_thms = inj_thms, (*FIXME: add real*)
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    66
    lessD = lessD,  (*We don't change LA_Data_Ref.lessD because the real ordering is dense!*)
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    67
    simpset = simpset addsimps add_rules
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    68
                      addsimprocs simprocs}),
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    69
  arith_discrete ("RealDef.real",false),
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    70
  Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]];
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    71
10760
02d727c441bb *** empty log message ***
nipkow
parents: 10758
diff changeset
    72
(* some thms for injection nat => real:
02d727c441bb *** empty log message ***
nipkow
parents: 10758
diff changeset
    73
real_of_nat_zero
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    74
?zero_eq_numeral_0
10760
02d727c441bb *** empty log message ***
nipkow
parents: 10758
diff changeset
    75
real_of_nat_add
02d727c441bb *** empty log message ***
nipkow
parents: 10758
diff changeset
    76
*)
02d727c441bb *** empty log message ***
nipkow
parents: 10758
diff changeset
    77
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    78
end;
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    79
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    80
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    81
(* Some test data [omitting examples that assume the ordering to be discrete!]
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    82
Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    83
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    84
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    85
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    86
Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    87
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    88
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    89
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    90
Goal "!!a::real. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    91
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    92
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    93
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    94
Goal "!!a::real. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    95
by (arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    96
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    97
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    98
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
    99
\     ==> a <= l";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   100
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   101
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   102
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   103
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   104
\     ==> a+a+a+a <= l+l+l+l";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   105
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   106
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   107
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   108
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   109
\     ==> a+a+a+a+a <= l+l+l+l+i";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   110
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   111
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   112
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   113
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   114
\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   115
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   116
qed "";
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   117
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   118
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   119
\     ==> 6*a <= 5*l+i";
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   120
by (fast_arith_tac 1);
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   121
qed "";
13554
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   122
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   123
Goal "a<=b ==> a < b+(1::real)";
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   124
by (fast_arith_tac 1);
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   125
qed "";
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   126
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   127
Goal "a<=b ==> a-(3::real) < b";
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   128
by (fast_arith_tac 1);
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   129
qed "";
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   130
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   131
Goal "a<=b ==> a-(1::real) < b";
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   132
by (fast_arith_tac 1);
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   133
qed "";
4679359bb218 Added missing rewrite rule and some arith examples
paulson
parents: 13462
diff changeset
   134
10722
55c8367bab05 rational linear arithmetic
nipkow
parents:
diff changeset
   135
*)