src/HOL/Real/RealBin.ML
author paulson
Thu, 01 Jan 2004 10:06:32 +0100
changeset 14334 6137d24eef79
parent 14329 ff3210fe968f
child 14352 a8b1a44d8264
permissions -rw-r--r--
tweaking of lemmas in RealDef, RealOrd
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
     1
(*  Title:      HOL/Real/RealBin.ML
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     2
    ID:         $Id$
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     5
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
     6
Binary arithmetic for the reals (integer literals only).
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     7
*)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     8
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10890
diff changeset
     9
(** real (coercion from int to real) **)
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    10
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10890
diff changeset
    11
Goal "real (number_of w :: int) = number_of w";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    12
by (simp_tac (simpset() addsimps [real_number_of_def]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    13
qed "real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    14
Addsimps [real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    15
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    16
Goalw [real_number_of_def] "Numeral0 = (0::real)";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    17
by (simp_tac (simpset() addsimps [real_of_int_zero RS sym]) 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    18
qed "real_numeral_0_eq_0";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    19
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    20
Goalw [real_number_of_def] "Numeral1 = (1::real)";
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    21
by (stac (real_of_one RS sym) 1);
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    22
by (Simp_tac 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    23
qed "real_numeral_1_eq_1";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    24
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    25
(** Addition **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    26
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    27
Goal "(number_of v :: real) + number_of v' = number_of (bin_add v v')";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    28
by (simp_tac
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    29
    (HOL_ss addsimps [real_number_of_def,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    30
                                  real_of_int_add, number_of_add]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    31
qed "add_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    32
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    33
Addsimps [add_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    34
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    35
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    36
(** Subtraction **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    37
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    38
Goalw [real_number_of_def] "- (number_of w :: real) = number_of (bin_minus w)";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    39
by (simp_tac
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
    40
    (HOL_ss addsimps [number_of_minus, real_of_int_minus]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    41
qed "minus_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    42
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    43
Goalw [real_number_of_def]
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    44
   "(number_of v :: real) - number_of w = number_of (bin_add v (bin_minus w))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    45
by (simp_tac
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
    46
    (HOL_ss addsimps [diff_number_of_eq, real_of_int_diff]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    47
qed "diff_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    48
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    49
Addsimps [minus_real_number_of, diff_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    50
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    51
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    52
(** Multiplication **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    53
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    54
Goal "(number_of v :: real) * number_of v' = number_of (bin_mult v v')";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    55
by (simp_tac
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    56
    (HOL_ss addsimps [real_number_of_def, real_of_int_mult,
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    57
                      number_of_mult]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    58
qed "mult_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    59
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    60
Addsimps [mult_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    61
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    62
Goal "(2::real) = 1 + 1";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    63
by (simp_tac (simpset() addsimps [real_numeral_1_eq_1 RS sym]) 1);
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    64
val lemma = result();
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    65
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    66
(*For specialist use: NOT as default simprules*)
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    67
Goal "2 * z = (z+z::real)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    68
by (simp_tac (simpset () addsimps [lemma, left_distrib]) 1);
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    69
qed "real_mult_2";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    70
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    71
Goal "z * 2 = (z+z::real)";
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    72
by (stac real_mult_commute 1 THEN rtac real_mult_2 1);
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    73
qed "real_mult_2_right";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    74
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    75
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    76
(*** Comparisons ***)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    77
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    78
(** Equals (=) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    79
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    80
Goal "((number_of v :: real) = number_of v') = \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    81
\     iszero (number_of (bin_add v (bin_minus v')))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    82
by (simp_tac
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    83
    (HOL_ss addsimps [real_number_of_def,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    84
                      real_of_int_inject, eq_number_of_eq]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    85
qed "eq_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    86
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    87
Addsimps [eq_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    88
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    89
(** Less-than (<) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    90
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    91
(*"neg" is used in rewrite rules for binary comparisons*)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    92
Goal "((number_of v :: real) < number_of v') = \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    93
\     neg (number_of (bin_add v (bin_minus v')))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    94
by (simp_tac
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    95
    (HOL_ss addsimps [real_number_of_def, real_of_int_less_iff,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
    96
                                  less_number_of_eq_neg]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    97
qed "less_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    98
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    99
Addsimps [less_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   100
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   101
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   102
(** Less-than-or-equals (<=) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   103
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   104
Goal "(number_of x <= (number_of y::real)) = \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   105
\     (~ number_of y < (number_of x::real))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   106
by (rtac (linorder_not_less RS sym) 1);
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   107
qed "le_real_number_of_eq_not_less";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   108
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   109
Addsimps [le_real_number_of_eq_not_less];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   110
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11704
diff changeset
   111
(*** New versions of existing theorems involving 0, 1 ***)
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   112
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   113
Goal "- 1 = (-1::real)";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   114
by (simp_tac (simpset() addsimps [real_numeral_1_eq_1 RS sym]) 1);
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   115
qed "real_minus_1_eq_m1";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   116
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   117
Goal "-1 * z = -(z::real)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   118
by (simp_tac (simpset() addsimps [real_minus_1_eq_m1 RS sym]) 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   119
qed "real_mult_minus1";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   120
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   121
Goal "z * -1 = -(z::real)";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   122
by (stac real_mult_commute 1 THEN rtac real_mult_minus1 1);
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   123
qed "real_mult_minus1_right";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   124
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   125
Addsimps [real_mult_minus1, real_mult_minus1_right];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   126
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   127
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11704
diff changeset
   128
(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   129
val real_numeral_ss =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   130
    HOL_ss addsimps [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   131
                     real_minus_1_eq_m1];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   132
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   133
fun rename_numerals th =
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
   134
    asm_full_simplify real_numeral_ss (Thm.transfer (the_context ()) th);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   135
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   136
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10890
diff changeset
   137
(** real from type "nat" **)
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   138
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   139
Goal "(0 < real (n::nat)) = (0<n)";
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   140
by (simp_tac (HOL_ss addsimps [real_of_nat_less_iff,
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   141
                               real_of_nat_zero RS sym]) 1);
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10752
diff changeset
   142
qed "zero_less_real_of_nat_iff";
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10752
diff changeset
   143
AddIffs [zero_less_real_of_nat_iff];
7583
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   144
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   145
Goal "(0 <= real (n::nat)) = (0<=n)";
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   146
by (simp_tac (HOL_ss addsimps [real_of_nat_le_iff,
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   147
                               real_of_nat_zero RS sym]) 1);
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10752
diff changeset
   148
qed "zero_le_real_of_nat_iff";
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10752
diff changeset
   149
AddIffs [zero_le_real_of_nat_iff];
7583
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   150
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   151
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   152
(*Like the ones above, for "equals"*)
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   153
AddIffs [real_of_nat_zero_iff];
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   154
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   155
(** Simplification of arithmetic when nested to the right **)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   156
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   157
Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::real)";
14290
84fda1b39947 removal of abel_cancel from Real
paulson
parents: 14270
diff changeset
   158
by (asm_full_simp_tac (simpset() addsimps [add_assoc RS sym]) 1); 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   159
qed "real_add_number_of_left";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   160
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   161
Goal "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::real)";
14290
84fda1b39947 removal of abel_cancel from Real
paulson
parents: 14270
diff changeset
   162
by (simp_tac (simpset() addsimps [mult_assoc RS sym]) 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   163
qed "real_mult_number_of_left";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   164
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   165
Goalw [real_diff_def]
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   166
     "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::real)";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   167
by (rtac real_add_number_of_left 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   168
qed "real_add_number_of_diff1";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   169
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   170
Goal "number_of v + (c - number_of w) = \
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   171
\     number_of (bin_add v (bin_minus w)) + (c::real)";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   172
by (stac (diff_real_number_of RS sym) 1);
14290
84fda1b39947 removal of abel_cancel from Real
paulson
parents: 14270
diff changeset
   173
by (asm_full_simp_tac (HOL_ss addsimps [real_diff_def]@add_ac) 1); 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   174
qed "real_add_number_of_diff2";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   175
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   176
Addsimps [real_add_number_of_left, real_mult_number_of_left,
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   177
          real_add_number_of_diff1, real_add_number_of_diff2];
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   178
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   179
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   180
(*"neg" is used in rewrite rules for binary comparisons*)
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10890
diff changeset
   181
Goal "real (number_of v :: nat) = \
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   182
\       (if neg (number_of v) then 0 \
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   183
\        else (number_of v :: real))";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   184
by (simp_tac
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
   185
    (HOL_ss addsimps [nat_number_of_def, real_of_nat_real_of_int,
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   186
                      real_of_nat_neg_int, real_number_of,
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   187
                      real_numeral_0_eq_0 RS sym]) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   188
qed "real_of_nat_number_of";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   189
Addsimps [real_of_nat_number_of];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   190
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   191
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   192
(**** Simprocs for numeric literals ****)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   193
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   194
(** For combine_numerals **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   195
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   196
Goal "i*u + (j*u + k) = (i+j)*u + (k::real)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   197
by (asm_simp_tac (simpset() addsimps [left_distrib] @ add_ac) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   198
qed "left_real_add_mult_distrib";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   199
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   200
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   201
(** For cancel_numerals **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   202
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   203
val rel_iff_rel_0_rls = map (inst "b" "?u+?v")
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   204
                   [less_iff_diff_less_0, eq_iff_diff_eq_0, le_iff_diff_le_0] @
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   205
                 map (inst "b" "n")
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   206
                   [less_iff_diff_less_0, eq_iff_diff_eq_0, le_iff_diff_le_0];
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   207
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   208
Goal "!!i::real. (i*u + m = j*u + n) = ((i-j)*u + m = n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   209
by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   210
                                     add_ac@rel_iff_rel_0_rls) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   211
qed "real_eq_add_iff1";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   212
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   213
Goal "!!i::real. (i*u + m = j*u + n) = (m = (j-i)*u + n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   214
by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   215
                                     add_ac@rel_iff_rel_0_rls) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   216
qed "real_eq_add_iff2";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   217
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   218
Goal "!!i::real. (i*u + m < j*u + n) = ((i-j)*u + m < n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   219
by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   220
                                     add_ac@rel_iff_rel_0_rls) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   221
qed "real_less_add_iff1";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   222
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   223
Goal "!!i::real. (i*u + m < j*u + n) = (m < (j-i)*u + n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   224
by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   225
                                     add_ac@rel_iff_rel_0_rls) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   226
qed "real_less_add_iff2";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   227
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   228
Goal "!!i::real. (i*u + m <= j*u + n) = ((i-j)*u + m <= n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   229
by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]@
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   230
                                     add_ac@rel_iff_rel_0_rls) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   231
qed "real_le_add_iff1";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   232
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   233
Goal "!!i::real. (i*u + m <= j*u + n) = (m <= (j-i)*u + n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   234
by (asm_simp_tac (simpset() addsimps [real_diff_def, left_distrib]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   235
                                     @add_ac@rel_iff_rel_0_rls) 1);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   236
qed "real_le_add_iff2";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   237
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   238
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   239
Addsimps [real_numeral_0_eq_0, real_numeral_1_eq_1];
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   240
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   241
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   242
structure Real_Numeral_Simprocs =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   243
struct
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   244
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   245
(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   246
  isn't complicated by the abstract 0 and 1.*)
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   247
val numeral_syms = [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym];
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   248
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   249
(*Utilities*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   250
10693
9e4a0e84d0d6 moved mk_bin from Numerals to HOLogic
nipkow
parents: 10689
diff changeset
   251
fun mk_numeral n = HOLogic.number_of_const HOLogic.realT $ HOLogic.mk_bin n;
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   252
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   253
(*Decodes a binary real constant, or 0, 1*)
12819
2f61bca07de7 slight re-use of code
paulson
parents: 12613
diff changeset
   254
val dest_numeral = Int_Numeral_Simprocs.dest_numeral;
2f61bca07de7 slight re-use of code
paulson
parents: 12613
diff changeset
   255
val find_first_numeral = Int_Numeral_Simprocs.find_first_numeral;
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   256
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   257
val zero = mk_numeral 0;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   258
val mk_plus = HOLogic.mk_binop "op +";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   259
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   260
val uminus_const = Const ("uminus", HOLogic.realT --> HOLogic.realT);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   261
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   262
(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   263
fun mk_sum []        = zero
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   264
  | mk_sum [t,u]     = mk_plus (t, u)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   265
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   266
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   267
(*this version ALWAYS includes a trailing zero*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   268
fun long_mk_sum []        = zero
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   269
  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   270
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   271
val dest_plus = HOLogic.dest_bin "op +" HOLogic.realT;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   272
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   273
(*decompose additions AND subtractions as a sum*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   274
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   275
        dest_summing (pos, t, dest_summing (pos, u, ts))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   276
  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   277
        dest_summing (pos, t, dest_summing (not pos, u, ts))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   278
  | dest_summing (pos, t, ts) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   279
        if pos then t::ts else uminus_const$t :: ts;
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   280
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   281
fun dest_sum t = dest_summing (true, t, []);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   282
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   283
val mk_diff = HOLogic.mk_binop "op -";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   284
val dest_diff = HOLogic.dest_bin "op -" HOLogic.realT;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   285
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   286
val one = mk_numeral 1;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   287
val mk_times = HOLogic.mk_binop "op *";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   288
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   289
fun mk_prod [] = one
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   290
  | mk_prod [t] = t
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   291
  | mk_prod (t :: ts) = if t = one then mk_prod ts
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   292
                        else mk_times (t, mk_prod ts);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   293
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   294
val dest_times = HOLogic.dest_bin "op *" HOLogic.realT;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   295
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   296
fun dest_prod t =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   297
      let val (t,u) = dest_times t
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   298
      in  dest_prod t @ dest_prod u  end
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   299
      handle TERM _ => [t];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   300
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   301
(*DON'T do the obvious simplifications; that would create special cases*)
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   302
fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   303
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   304
(*Express t as a product of (possibly) a numeral with other sorted terms*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   305
fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   306
  | dest_coeff sign t =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   307
    let val ts = sort Term.term_ord (dest_prod t)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   308
        val (n, ts') = find_first_numeral [] ts
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   309
                          handle TERM _ => (1, ts)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   310
    in (sign*n, mk_prod ts') end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   311
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   312
(*Find first coefficient-term THAT MATCHES u*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   313
fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   314
  | find_first_coeff past u (t::terms) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   315
        let val (n,u') = dest_coeff 1 t
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   316
        in  if u aconv u' then (n, rev past @ terms)
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   317
                          else find_first_coeff (t::past) u terms
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   318
        end
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   319
        handle TERM _ => find_first_coeff (t::past) u terms;
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   320
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   321
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   322
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   323
val add_0s  = map rename_numerals [real_add_zero_left, real_add_zero_right];
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   324
val mult_1s = map rename_numerals [real_mult_1, real_mult_1_right] @
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   325
              [real_mult_minus1, real_mult_minus1_right];
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   326
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   327
(*To perform binary arithmetic*)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9080
diff changeset
   328
val bin_simps =
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   329
    [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   330
     add_real_number_of, real_add_number_of_left, minus_real_number_of,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   331
     diff_real_number_of, mult_real_number_of, real_mult_number_of_left] @
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   332
    bin_arith_simps @ bin_rel_simps;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   333
14123
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   334
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   335
  during re-arrangement*)
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   336
val non_add_bin_simps = 
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   337
    bin_simps \\ [real_add_number_of_left, add_real_number_of];
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   338
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   339
(*To evaluate binary negations of coefficients*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   340
val real_minus_simps = NCons_simps @
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   341
                   [real_minus_1_eq_m1, minus_real_number_of,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   342
                    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   343
                    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   344
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   345
(*To let us treat subtraction as addition*)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   346
val diff_simps = [real_diff_def, minus_add_distrib, minus_minus];
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   347
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   348
(*to extract again any uncancelled minuses*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   349
val real_minus_from_mult_simps =
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   350
    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym];
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   351
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   352
(*combine unary minus with numeric literals, however nested within a product*)
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   353
val real_mult_minus_simps =
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   354
    [real_mult_assoc, minus_mult_left, real_minus_mult_commute];
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   355
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   356
(*Apply the given rewrite (if present) just once*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   357
fun trans_tac None      = all_tac
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   358
  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   359
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   360
(*Final simplification: cancel + and *  *)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   361
val simplify_meta_eq =
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   362
    Int_Numeral_Simprocs.simplify_meta_eq
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   363
         [add_0, add_0_right,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   364
          mult_left_zero, mult_right_zero, mult_1, mult_1_right];
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   365
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   366
fun prep_simproc (name, pats, proc) =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   367
  Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   368
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   369
structure CancelNumeralsCommon =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   370
  struct
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   371
  val mk_sum            = mk_sum
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   372
  val dest_sum          = dest_sum
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   373
  val mk_coeff          = mk_coeff
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   374
  val dest_coeff        = dest_coeff 1
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   375
  val find_first_coeff  = find_first_coeff []
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   376
  val trans_tac         = trans_tac
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   377
  val norm_tac =
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   378
     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   379
                                         real_minus_simps@add_ac))
14123
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   380
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   381
     THEN ALLGOALS
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   382
              (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   383
                                         add_ac@mult_ac))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   384
  val numeral_simp_tac  = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   385
  val simplify_meta_eq  = simplify_meta_eq
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   386
  end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   387
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   388
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   389
structure EqCancelNumerals = CancelNumeralsFun
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   390
 (open CancelNumeralsCommon
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   391
  val prove_conv = Bin_Simprocs.prove_conv
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   392
  val mk_bal   = HOLogic.mk_eq
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   393
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   394
  val bal_add1 = real_eq_add_iff1 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   395
  val bal_add2 = real_eq_add_iff2 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   396
);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   397
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   398
structure LessCancelNumerals = CancelNumeralsFun
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   399
 (open CancelNumeralsCommon
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   400
  val prove_conv = Bin_Simprocs.prove_conv
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   401
  val mk_bal   = HOLogic.mk_binrel "op <"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   402
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   403
  val bal_add1 = real_less_add_iff1 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   404
  val bal_add2 = real_less_add_iff2 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   405
);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   406
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   407
structure LeCancelNumerals = CancelNumeralsFun
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   408
 (open CancelNumeralsCommon
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   409
  val prove_conv = Bin_Simprocs.prove_conv
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   410
  val mk_bal   = HOLogic.mk_binrel "op <="
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   411
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   412
  val bal_add1 = real_le_add_iff1 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   413
  val bal_add2 = real_le_add_iff2 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   414
);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   415
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   416
val cancel_numerals =
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   417
  map prep_simproc
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   418
   [("realeq_cancel_numerals",
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   419
     ["(l::real) + m = n", "(l::real) = m + n",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   420
      "(l::real) - m = n", "(l::real) = m - n",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   421
      "(l::real) * m = n", "(l::real) = m * n"],
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   422
     EqCancelNumerals.proc),
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   423
    ("realless_cancel_numerals",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   424
     ["(l::real) + m < n", "(l::real) < m + n",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   425
      "(l::real) - m < n", "(l::real) < m - n",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   426
      "(l::real) * m < n", "(l::real) < m * n"],
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   427
     LessCancelNumerals.proc),
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   428
    ("realle_cancel_numerals",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   429
     ["(l::real) + m <= n", "(l::real) <= m + n",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   430
      "(l::real) - m <= n", "(l::real) <= m - n",
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   431
      "(l::real) * m <= n", "(l::real) <= m * n"],
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   432
     LeCancelNumerals.proc)];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   433
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   434
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   435
structure CombineNumeralsData =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   436
  struct
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   437
  val add               = op + : int*int -> int
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   438
  val mk_sum            = long_mk_sum    (*to work for e.g. 2*x + 3*x *)
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   439
  val dest_sum          = dest_sum
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   440
  val mk_coeff          = mk_coeff
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   441
  val dest_coeff        = dest_coeff 1
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   442
  val left_distrib      = left_real_add_mult_distrib RS trans
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   443
  val prove_conv        = Bin_Simprocs.prove_conv_nohyps
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   444
  val trans_tac         = trans_tac
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   445
  val norm_tac =
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   446
     ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   447
                                   diff_simps@real_minus_simps@add_ac))
14123
b300efca4ae0 fixed simprocs
paulson
parents: 13480
diff changeset
   448
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@real_mult_minus_simps))
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   449
     THEN ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   450
                                              add_ac@mult_ac))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   451
  val numeral_simp_tac  = ALLGOALS
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   452
                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   453
  val simplify_meta_eq  =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   454
        Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   455
  end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   456
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   457
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   458
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   459
val combine_numerals =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   460
  prep_simproc ("real_combine_numerals", ["(i::real) + j", "(i::real) - j"], CombineNumerals.proc);
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   461
10689
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   462
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   463
(** Declarations for ExtractCommonTerm **)
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   464
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   465
(*this version ALWAYS includes a trailing one*)
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   466
fun long_mk_prod []        = one
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   467
  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   468
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   469
(*Find first term that matches u*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   470
fun find_first past u []         = raise TERM("find_first", [])
10689
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   471
  | find_first past u (t::terms) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   472
        if u aconv t then (rev past @ terms)
10689
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   473
        else find_first (t::past) u terms
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   474
        handle TERM _ => find_first (t::past) u terms;
10689
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   475
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   476
(*Final simplification: cancel + and *  *)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   477
fun cancel_simplify_meta_eq cancel_th th =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   478
    Int_Numeral_Simprocs.simplify_meta_eq
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   479
        [real_mult_1, real_mult_1_right]
10689
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   480
        (([th, cancel_th]) MRS trans);
5c44de6aadf4 loads the new simproc extract_common_term
paulson
parents: 10648
diff changeset
   481
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   482
(*** Making constant folding work for 0 and 1 too ***)
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   483
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   484
structure RealAbstractNumeralsData =
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   485
  struct
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   486
  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   487
  val is_numeral      = Bin_Simprocs.is_numeral
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   488
  val numeral_0_eq_0  = real_numeral_0_eq_0
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   489
  val numeral_1_eq_1  = real_numeral_1_eq_1
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   490
  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   491
  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   492
  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   493
  end
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   494
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   495
structure RealAbstractNumerals = AbstractNumeralsFun (RealAbstractNumeralsData)
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   496
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   497
(*For addition, we already have rules for the operand 0.
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   498
  Multiplication is omitted because there are already special rules for
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   499
  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   500
  For the others, having three patterns is a compromise between just having
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   501
  one (many spurious calls) and having nine (just too many!) *)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   502
val eval_numerals =
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   503
  map prep_simproc
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   504
   [("real_add_eval_numerals",
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   505
     ["(m::real) + 1", "(m::real) + number_of v"],
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   506
     RealAbstractNumerals.proc add_real_number_of),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   507
    ("real_diff_eval_numerals",
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   508
     ["(m::real) - 1", "(m::real) - number_of v"],
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   509
     RealAbstractNumerals.proc diff_real_number_of),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   510
    ("real_eq_eval_numerals",
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   511
     ["(m::real) = 0", "(m::real) = 1", "(m::real) = number_of v"],
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   512
     RealAbstractNumerals.proc eq_real_number_of),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   513
    ("real_less_eval_numerals",
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   514
     ["(m::real) < 0", "(m::real) < 1", "(m::real) < number_of v"],
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   515
     RealAbstractNumerals.proc less_real_number_of),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   516
    ("real_le_eval_numerals",
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   517
     ["(m::real) <= 0", "(m::real) <= 1", "(m::real) <= number_of v"],
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   518
     RealAbstractNumerals.proc le_real_number_of_eq_not_less)]
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   519
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   520
end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   521
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   522
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   523
Addsimprocs Real_Numeral_Simprocs.eval_numerals;
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   524
Addsimprocs Real_Numeral_Simprocs.cancel_numerals;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   525
Addsimprocs [Real_Numeral_Simprocs.combine_numerals];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   526
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   527
(*examples:
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   528
print_depth 22;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   529
set timing;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   530
set trace_simp;
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   531
fun test s = (Goal s; by (Simp_tac 1));
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   532
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   533
test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::real)";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   534
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   535
test "2*u = (u::real)";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   536
test "(i + j + 12 + (k::real)) - 15 = y";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   537
test "(i + j + 12 + (k::real)) - 5 = y";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   538
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   539
test "y - b < (b::real)";
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   540
test "y - (3*b + c) < (b::real) - 2*c";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   541
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   542
test "(2*x - (u*v) + y) - v*3*u = (w::real)";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   543
test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::real)";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   544
test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::real)";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   545
test "u*v - (x*u*v + (u*v)*4 + y) = (w::real)";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   546
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   547
test "(i + j + 12 + (k::real)) = u + 15 + y";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   548
test "(i + j*2 + 12 + (k::real)) = j + 5 + y";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   549
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   550
test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::real)";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   551
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   552
test "a + -(b+c) + b = (d::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   553
test "a + -(b+c) - b = (d::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   554
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   555
(*negative numerals*)
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   556
test "(i + j + -2 + (k::real)) - (u + 5 + y) = zz";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   557
test "(i + j + -3 + (k::real)) < u + 5 + y";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   558
test "(i + j + 3 + (k::real)) < u + -6 + y";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   559
test "(i + j + -12 + (k::real)) - 15 = y";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   560
test "(i + j + 12 + (k::real)) - -15 = y";
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   561
test "(i + j + -12 + (k::real)) - -15 = y";
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   562
*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   563
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   564
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   565
(** Constant folding for real plus and times **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   566
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   567
(*We do not need
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   568
    structure Real_Plus_Assoc = Assoc_Fold (Real_Plus_Assoc_Data);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   569
  because combine_numerals does the same thing*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   570
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   571
structure Real_Times_Assoc_Data : ASSOC_FOLD_DATA =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   572
struct
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   573
  val ss                = HOL_ss
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   574
  val eq_reflection     = eq_reflection
9433
ac20534ce4d1 avoid referencing thy value;
wenzelm
parents: 9108
diff changeset
   575
  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12819
diff changeset
   576
  val T      = HOLogic.realT
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   577
  val plus   = Const ("op *", [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   578
  val add_ac = mult_ac
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   579
end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   580
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   581
structure Real_Times_Assoc = Assoc_Fold (Real_Times_Assoc_Data);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   582
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   583
Addsimprocs [Real_Times_Assoc.conv];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   584
10699
f0c3da8477e9 more tidying
paulson
parents: 10693
diff changeset
   585
(*Simplification of  x-y < 0, etc.*)
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   586
AddIffs [less_iff_diff_less_0 RS sym];
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   587
AddIffs [eq_iff_diff_eq_0 RS sym];
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14290
diff changeset
   588
AddIffs [le_iff_diff_le_0 RS sym];
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10752
diff changeset
   589
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
   590
Addsimps [real_minus_1_eq_m1];