src/HOL/Real/RealBin.ML
author paulson
Wed, 14 Jun 2000 18:19:20 +0200
changeset 9068 202fdf72cf23
parent 9043 ca761fe227d8
child 9080 67ca888af420
permissions -rw-r--r--
installing the cancel_numerals and combine_numerals simprocs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     1
(*  Title:      HOL/RealBin.ML
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
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     6
Binary arithmetic for the reals (integer literals only)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     7
*)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     8
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
     9
(** real_of_int (coercion from int to real) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    10
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    11
Goal "real_of_int (number_of w) = number_of w";
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
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    16
Goalw [real_number_of_def] "(0::real) = #0";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    17
by (simp_tac (simpset() addsimps [real_of_int_zero RS sym]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    18
qed "zero_eq_numeral_0";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    19
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    20
Goalw [real_number_of_def] "1r = #1";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    21
by (simp_tac (simpset() addsimps [real_of_int_one RS sym]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    22
qed "one_eq_numeral_1";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    23
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    24
(** Addition **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    25
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    26
Goal "(number_of v :: real) + number_of v' = number_of (bin_add v v')";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    27
by (simp_tac
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    28
    (simpset_of Int.thy addsimps [real_number_of_def, 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    29
				  real_of_int_add, number_of_add]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    30
qed "add_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    31
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    32
Addsimps [add_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    33
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    34
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    35
(** Subtraction **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    36
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    37
Goalw [real_number_of_def] "- (number_of w :: real) = number_of (bin_minus w)";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    38
by (simp_tac
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    39
    (simpset_of Int.thy addsimps [number_of_minus, real_of_int_minus]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    40
qed "minus_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    41
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    42
Goalw [real_number_of_def]
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    43
   "(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
    44
by (simp_tac
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    45
    (simpset_of Int.thy addsimps [diff_number_of_eq, real_of_int_diff]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    46
qed "diff_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    47
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    48
Addsimps [minus_real_number_of, diff_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    49
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    50
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    51
(** Multiplication **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    52
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    53
Goal "(number_of v :: real) * number_of v' = number_of (bin_mult v v')";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    54
by (simp_tac
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    55
    (simpset_of Int.thy addsimps [real_number_of_def, 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    56
				  real_of_int_mult, number_of_mult]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    57
qed "mult_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    58
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    59
Addsimps [mult_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    60
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    61
Goal "(#2::real) = #1 + #1";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    62
by (Simp_tac 1);
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    63
val lemma = result();
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    64
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    65
(*For specialist use: NOT as default simprules*)
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    66
Goal "#2 * z = (z+z::real)";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    67
by (simp_tac (simpset_of RealDef.thy
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    68
	      addsimps [lemma, real_add_mult_distrib,
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    69
			one_eq_numeral_1 RS sym]) 1);
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    70
qed "real_mult_2";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    71
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    72
Goal "z * #2 = (z+z::real)";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    73
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
    74
qed "real_mult_2_right";
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7583
diff changeset
    75
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    76
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    77
(*** Comparisons ***)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    78
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    79
(** Equals (=) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    80
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    81
Goal "((number_of v :: real) = number_of v') = \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    82
\     iszero (number_of (bin_add v (bin_minus v')))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    83
by (simp_tac
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    84
    (simpset_of Int.thy addsimps [real_number_of_def, 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    85
				  real_of_int_eq_iff, eq_number_of_eq]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    86
qed "eq_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    87
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    88
Addsimps [eq_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    89
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    90
(** Less-than (<) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    91
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    92
(*"neg" is used in rewrite rules for binary comparisons*)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    93
Goal "((number_of v :: real) < number_of v') = \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    94
\     neg (number_of (bin_add v (bin_minus v')))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    95
by (simp_tac
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    96
    (simpset_of Int.thy addsimps [real_number_of_def, real_of_int_less_iff, 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    97
				  less_number_of_eq_neg]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    98
qed "less_real_number_of";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    99
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   100
Addsimps [less_real_number_of];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   101
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   102
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   103
(** Less-than-or-equals (<=) **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   104
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   105
Goal "(number_of x <= (number_of y::real)) = \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   106
\     (~ number_of y < (number_of x::real))";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   107
by (rtac (linorder_not_less RS sym) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   108
qed "le_real_number_of_eq_not_less"; 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   109
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   110
Addsimps [le_real_number_of_eq_not_less];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   111
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   112
(*** New versions of existing theorems involving 0, 1r ***)
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   113
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   114
Goal "- #1 = (#-1::real)";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   115
by (Simp_tac 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   116
qed "minus_numeral_one";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   117
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   118
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   119
(*Maps 0 to #0 and 1r to #1 and -1r to #-1*)
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   120
val real_numeral_ss = 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   121
    HOL_ss addsimps [zero_eq_numeral_0, one_eq_numeral_1, 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   122
		     minus_numeral_one];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   123
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   124
fun rename_numerals thy th = 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   125
    asm_full_simplify real_numeral_ss (change_theory thy th);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   126
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   127
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   128
(*Now insert some identities previously stated for 0 and 1r*)
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   129
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   130
(** RealDef & Real **)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   131
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   132
Addsimps (map (rename_numerals thy) 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   133
	  [real_minus_zero, real_minus_zero_iff,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   134
	   real_add_zero_left, real_add_zero_right, 
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   135
	   real_diff_0, real_diff_0_right,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   136
	   real_mult_0_right, real_mult_0, real_mult_1_right, real_mult_1,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   137
	   real_mult_minus_1_right, real_mult_minus_1, real_rinv_1,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   138
	   real_minus_zero_less_iff]);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   139
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   140
AddIffs (map (rename_numerals thy) [real_mult_is_0, real_0_is_mult]);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   141
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   142
bind_thm ("real_0_less_mult_iff", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   143
	  rename_numerals thy real_zero_less_mult_iff);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   144
bind_thm ("real_0_le_mult_iff", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   145
	  rename_numerals thy real_zero_le_mult_iff);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   146
bind_thm ("real_mult_less_0_iff", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   147
	  rename_numerals thy real_mult_less_zero_iff);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   148
bind_thm ("real_mult_le_0_iff", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   149
	  rename_numerals thy real_mult_le_zero_iff);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   150
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   151
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   152
(*Perhaps add some theorems that aren't in the default simpset, as
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   153
  done in Integ/NatBin.ML*)
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   154
7583
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   155
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   156
(*  Author:     Tobias Nipkow, TU Muenchen
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   157
    Copyright   1999 TU Muenchen
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   158
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   159
Instantiate linear arithmetic decision procedure for the reals.
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   160
FIXME: multiplication with constants (eg #2 * x) does not work yet.
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   161
Solution: there should be a simproc for combining coefficients.
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   162
*)
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   163
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   164
let
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   165
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   166
(* reduce contradictory <= to False *)
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   167
val simps = [order_less_irrefl,zero_eq_numeral_0,one_eq_numeral_1,
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   168
  add_real_number_of,minus_real_number_of,diff_real_number_of,
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   169
  mult_real_number_of,eq_real_number_of,less_real_number_of,
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   170
  le_real_number_of_eq_not_less];
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   171
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   172
val simprocs = [Real_Cancel.sum_conv, Real_Cancel.rel_conv];
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   173
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   174
val add_mono_thms =
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   175
  map (fn s => prove_goal thy s
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   176
                 (fn prems => [cut_facts_tac prems 1,
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   177
                      asm_simp_tac (simpset() addsimps
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   178
     [real_add_le_mono,real_add_less_mono,
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   179
      real_add_less_le_mono,real_add_le_less_mono]) 1]))
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   180
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   181
     "(i  = j) & (k <= l) ==> i + k <= j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   182
     "(i <= j) & (k  = l) ==> i + k <= j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   183
     "(i  = j) & (k  = l) ==> i + k  = j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   184
     "(i < j) & (k = l)   ==> i + k < j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   185
     "(i = j) & (k < l)   ==> i + k < j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   186
     "(i < j) & (k <= l)  ==> i + k < j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   187
     "(i <= j) & (k < l)  ==> i + k < j + (l::real)",
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   188
     "(i < j) & (k < l)   ==> i + k < j + (l::real)"];
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   189
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   190
in
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   191
LA_Data_Ref.add_mono_thms := !LA_Data_Ref.add_mono_thms @ add_mono_thms;
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   192
LA_Data_Ref.ss_ref := !LA_Data_Ref.ss_ref addsimps simps
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   193
                      addsimprocs simprocs;
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   194
LA_Data_Ref.discrete := !LA_Data_Ref.discrete @ [("RealDef.real",false)]
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   195
end;
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   196
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   197
let
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   198
val real_arith_simproc_pats =
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   199
  map (fn s => Thm.read_cterm (Theory.sign_of thy) (s, HOLogic.boolT))
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   200
      ["(m::real) < n","(m::real) <= n", "(m::real) = n"];
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   201
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   202
val fast_real_arith_simproc = mk_simproc
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   203
  "fast_real_arith" real_arith_simproc_pats Fast_Arith.lin_arith_prover;
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   204
in
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   205
Addsimprocs [fast_real_arith_simproc]
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   206
end;
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   207
 
7583
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   208
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   209
Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
7583
d1b40e0464b1 Restructured lin.arith.package and fixed a proof in RComplete.
nipkow
parents: 7292
diff changeset
   210
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   211
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   212
(** Simplification of arithmetic when nested to the right **)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   213
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   214
Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::real)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   215
by Auto_tac; 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   216
qed "real_add_number_of_left";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   217
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   218
Goal "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::real)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   219
by (simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   220
qed "real_mult_number_of_left";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   221
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   222
Goalw [real_diff_def]
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   223
    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::real)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   224
by (rtac real_add_number_of_left 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   225
qed "real_add_number_of_diff1";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   226
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   227
Goal "number_of v + (c - number_of w) = \
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   228
\    number_of (bin_add v (bin_minus w)) + (c::real)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   229
by (stac (diff_real_number_of RS sym) 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   230
by Auto_tac;
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   231
qed "real_add_number_of_diff2";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   232
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   233
Addsimps [real_add_number_of_left, real_mult_number_of_left,
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   234
	  real_add_number_of_diff1, real_add_number_of_diff2]; 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   235
9068
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   236
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   237
(*"neg" is used in rewrite rules for binary comparisons*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   238
Goal "real_of_nat (number_of v :: nat) = \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   239
\        (if neg (number_of v) then #0 \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   240
\         else (number_of v :: real))";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   241
by (simp_tac
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   242
    (simpset_of Int.thy addsimps [nat_number_of_def, real_of_nat_real_of_int,
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   243
				  real_of_nat_neg_int, real_number_of,
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   244
				  zero_eq_numeral_0]) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   245
qed "real_of_nat_number_of";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   246
Addsimps [real_of_nat_number_of];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   247
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   248
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   249
(**** Simprocs for numeric literals ****)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   250
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   251
(** Combining of literal coefficients in sums of products **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   252
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   253
Goal "(x < y) = (x-y < (#0::real))";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   254
by Auto_tac; 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   255
qed "real_less_iff_diff_less_0";
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
Goal "(x = y) = (x-y = (#0::real))";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   258
by Auto_tac; 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   259
qed "real_eq_iff_diff_eq_0";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   260
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   261
Goal "(x <= y) = (x-y <= (#0::real))";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   262
by Auto_tac; 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   263
qed "real_le_iff_diff_le_0";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   264
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   265
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   266
(** For combine_numerals **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   267
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   268
Goal "i*u + (j*u + k) = (i+j)*u + (k::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   269
by (asm_simp_tac (simpset() addsimps [real_add_mult_distrib]) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   270
qed "left_real_add_mult_distrib";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   271
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
(** For cancel_numerals **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   274
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   275
val rel_iff_rel_0_rls = map (inst "y" "?u+?v")
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   276
                          [real_less_iff_diff_less_0, real_eq_iff_diff_eq_0, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   277
			   real_le_iff_diff_le_0] @
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   278
		        map (inst "y" "n")
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   279
                          [real_less_iff_diff_less_0, real_eq_iff_diff_eq_0, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   280
			   real_le_iff_diff_le_0];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   281
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   282
Goal "!!i::real. (i*u + m = j*u + n) = ((i-j)*u + m = n)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   283
by (asm_simp_tac (simpset() addsimps [real_diff_def, real_add_mult_distrib]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   284
		                     real_add_ac@rel_iff_rel_0_rls) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   285
qed "real_eq_add_iff1";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   286
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   287
Goal "!!i::real. (i*u + m = j*u + n) = (m = (j-i)*u + n)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   288
by (asm_simp_tac (simpset() addsimps [real_diff_def, real_add_mult_distrib]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   289
                                     real_add_ac@rel_iff_rel_0_rls) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   290
qed "real_eq_add_iff2";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   291
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   292
Goal "!!i::real. (i*u + m < j*u + n) = ((i-j)*u + m < n)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   293
by (asm_simp_tac (simpset() addsimps [real_diff_def, real_add_mult_distrib]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   294
                                     real_add_ac@rel_iff_rel_0_rls) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   295
qed "real_less_add_iff1";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   296
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   297
Goal "!!i::real. (i*u + m < j*u + n) = (m < (j-i)*u + n)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   298
by (asm_simp_tac (simpset() addsimps [real_diff_def, real_add_mult_distrib]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   299
                                     real_add_ac@rel_iff_rel_0_rls) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   300
qed "real_less_add_iff2";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   301
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   302
Goal "!!i::real. (i*u + m <= j*u + n) = ((i-j)*u + m <= n)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   303
by (asm_simp_tac (simpset() addsimps [real_diff_def, real_add_mult_distrib]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   304
                                     real_add_ac@rel_iff_rel_0_rls) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   305
qed "real_le_add_iff1";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   306
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   307
Goal "!!i::real. (i*u + m <= j*u + n) = (m <= (j-i)*u + n)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   308
by (asm_simp_tac (simpset() addsimps [real_diff_def, real_add_mult_distrib]
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   309
                                     @real_add_ac@rel_iff_rel_0_rls) 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   310
qed "real_le_add_iff2";
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
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   313
structure Real_Numeral_Simprocs =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   314
struct
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   315
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   316
(*Utilities*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   317
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   318
fun mk_numeral n = HOLogic.number_of_const HOLogic.realT $ 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   319
                   NumeralSyntax.mk_bin n;
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
(*Decodes a binary real constant*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   322
fun dest_numeral (Const("Numeral.number_of", _) $ w) = 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   323
     (NumeralSyntax.dest_bin w
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   324
      handle Match => raise TERM("Real_Numeral_Simprocs.dest_numeral:1", [w]))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   325
  | dest_numeral t = raise TERM("Real_Numeral_Simprocs.dest_numeral:2", [t]);
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
fun find_first_numeral past (t::terms) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   328
	((dest_numeral t, rev past @ terms)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   329
	 handle TERM _ => find_first_numeral (t::past) terms)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   330
  | find_first_numeral past [] = raise TERM("find_first_numeral", []);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   331
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   332
val zero = mk_numeral 0;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   333
val mk_plus = HOLogic.mk_binop "op +";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   334
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   335
val uminus_const = Const ("uminus", HOLogic.realT --> HOLogic.realT);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   336
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   337
(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   338
fun mk_sum []        = zero
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   339
  | mk_sum [t,u]     = mk_plus (t, u)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   340
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   341
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   342
(*this version ALWAYS includes a trailing zero*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   343
fun long_mk_sum []        = zero
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   344
  | 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
   345
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   346
val dest_plus = HOLogic.dest_bin "op +" HOLogic.realT;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   347
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   348
(*decompose additions AND subtractions as a sum*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   349
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   350
        dest_summing (pos, t, dest_summing (pos, u, ts))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   351
  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   352
        dest_summing (pos, t, dest_summing (not pos, u, ts))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   353
  | dest_summing (pos, t, ts) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   354
	if pos then t::ts else uminus_const$t :: ts;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   355
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   356
fun dest_sum t = dest_summing (true, t, []);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   357
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   358
val mk_diff = HOLogic.mk_binop "op -";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   359
val dest_diff = HOLogic.dest_bin "op -" HOLogic.realT;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   360
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   361
val one = mk_numeral 1;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   362
val mk_times = HOLogic.mk_binop "op *";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   363
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   364
fun mk_prod [] = one
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   365
  | mk_prod [t] = t
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   366
  | 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
   367
                        else mk_times (t, mk_prod ts);
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
val dest_times = HOLogic.dest_bin "op *" HOLogic.realT;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   370
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   371
fun dest_prod t =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   372
      let val (t,u) = dest_times t 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   373
      in  dest_prod t @ dest_prod u  end
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   374
      handle TERM _ => [t];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   375
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   376
(*DON'T do the obvious simplifications; that would create special cases*) 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   377
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
   378
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   379
(*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
   380
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
   381
  | dest_coeff sign t =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   382
    let val ts = sort Term.term_ord (dest_prod t)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   383
	val (n, ts') = find_first_numeral [] ts
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   384
                          handle TERM _ => (1, ts)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   385
    in (sign*n, mk_prod ts') end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   386
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   387
(*Find first coefficient-term THAT MATCHES u*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   388
fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   389
  | find_first_coeff past u (t::terms) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   390
	let val (n,u') = dest_coeff 1 t
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   391
	in  if u aconv u' then (n, rev past @ terms)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   392
			  else find_first_coeff (t::past) u terms
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   393
	end
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   394
	handle TERM _ => find_first_coeff (t::past) u terms;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   395
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
(*Simplify #1*n and n*#1 to n*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   398
val add_0s = map (rename_numerals thy) 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   399
                 [real_add_zero_left, real_add_zero_right];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   400
val mult_1s = map (rename_numerals thy) 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   401
                  [real_mult_1, real_mult_1_right, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   402
		   real_mult_minus_1, real_mult_minus_1_right];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   403
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   404
(*To perform binary arithmetic*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   405
val bin_simps = 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   406
    [add_real_number_of, real_add_number_of_left, minus_real_number_of, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   407
     diff_real_number_of] @ 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   408
    bin_arith_simps @ bin_rel_simps;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   409
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   410
(*To evaluate binary negations of coefficients*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   411
val real_minus_simps = NCons_simps @
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   412
                   [minus_real_number_of, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   413
		    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   414
		    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   415
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   416
(*To let us treat subtraction as addition*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   417
val diff_simps = [real_diff_def, real_minus_add_distrib, real_minus_minus];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   418
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   419
(*Apply the given rewrite (if present) just once*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   420
fun trans_tac None      = all_tac
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   421
  | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   422
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   423
fun prove_conv name tacs sg (t, u) =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   424
  if t aconv u then None
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   425
  else
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   426
  let val ct = cterm_of sg (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   427
  in Some
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   428
     (prove_goalw_cterm [] ct (K tacs)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   429
      handle ERROR => error 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   430
	  ("The error(s) above occurred while trying to prove " ^
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   431
	   string_of_cterm ct ^ "\nInternal failure of simproc " ^ name))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   432
  end;
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
(*Final simplification: cancel + and *  *)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   435
val simplify_meta_eq = 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   436
    Int_Numeral_Simprocs.simplify_meta_eq
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   437
         [real_add_zero_left, real_add_zero_right,
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   438
 	  real_mult_0, real_mult_0_right, real_mult_1, real_mult_1_right];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   439
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   440
fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   441
fun prep_pat s = Thm.read_cterm (Theory.sign_of RealInt.thy) (s, HOLogic.termT);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   442
val prep_pats = map prep_pat;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   443
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   444
structure CancelNumeralsCommon =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   445
  struct
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   446
  val mk_sum    	= mk_sum
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   447
  val dest_sum		= dest_sum
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   448
  val mk_coeff		= mk_coeff
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   449
  val dest_coeff	= dest_coeff 1
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   450
  val find_first_coeff	= find_first_coeff []
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   451
  val trans_tac         = trans_tac
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   452
  val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   453
                                                real_minus_simps@real_add_ac))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   454
                 THEN ALLGOALS
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   455
                    (simp_tac (HOL_ss addsimps [real_minus_mult_eq2]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   456
                                         bin_simps@real_add_ac@real_mult_ac))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   457
  val numeral_simp_tac	= ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   458
  val simplify_meta_eq  = simplify_meta_eq
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   459
  end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   460
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   461
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   462
structure EqCancelNumerals = CancelNumeralsFun
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   463
 (open CancelNumeralsCommon
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   464
  val prove_conv = prove_conv "realeq_cancel_numerals"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   465
  val mk_bal   = HOLogic.mk_eq
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   466
  val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   467
  val bal_add1 = real_eq_add_iff1 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   468
  val bal_add2 = real_eq_add_iff2 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   469
);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   470
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   471
structure LessCancelNumerals = CancelNumeralsFun
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   472
 (open CancelNumeralsCommon
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   473
  val prove_conv = prove_conv "realless_cancel_numerals"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   474
  val mk_bal   = HOLogic.mk_binrel "op <"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   475
  val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   476
  val bal_add1 = real_less_add_iff1 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   477
  val bal_add2 = real_less_add_iff2 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   478
);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   479
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   480
structure LeCancelNumerals = CancelNumeralsFun
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   481
 (open CancelNumeralsCommon
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   482
  val prove_conv = prove_conv "realle_cancel_numerals"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   483
  val mk_bal   = HOLogic.mk_binrel "op <="
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   484
  val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   485
  val bal_add1 = real_le_add_iff1 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   486
  val bal_add2 = real_le_add_iff2 RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   487
);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   488
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   489
val cancel_numerals = 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   490
  map prep_simproc
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   491
   [("realeq_cancel_numerals",
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   492
     prep_pats ["(l::real) + m = n", "(l::real) = m + n", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   493
		"(l::real) - m = n", "(l::real) = m - n", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   494
		"(l::real) * m = n", "(l::real) = m * n"], 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   495
     EqCancelNumerals.proc),
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   496
    ("realless_cancel_numerals", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   497
     prep_pats ["(l::real) + m < n", "(l::real) < m + n", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   498
		"(l::real) - m < n", "(l::real) < m - n", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   499
		"(l::real) * m < n", "(l::real) < m * n"], 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   500
     LessCancelNumerals.proc),
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   501
    ("realle_cancel_numerals", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   502
     prep_pats ["(l::real) + m <= n", "(l::real) <= m + n", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   503
		"(l::real) - m <= n", "(l::real) <= m - n", 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   504
		"(l::real) * m <= n", "(l::real) <= m * n"], 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   505
     LeCancelNumerals.proc)];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   506
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   507
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   508
structure CombineNumeralsData =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   509
  struct
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   510
  val mk_sum    	= long_mk_sum    (*to work for e.g. #2*x + #3*x *)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   511
  val dest_sum		= dest_sum
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   512
  val mk_coeff		= mk_coeff
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   513
  val dest_coeff	= dest_coeff 1
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   514
  val left_distrib	= left_real_add_mult_distrib RS trans
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   515
  val prove_conv	= prove_conv "real_combine_numerals"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   516
  val trans_tac          = trans_tac
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   517
  val norm_tac = ALLGOALS
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   518
                   (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   519
                                              real_minus_simps@real_add_ac))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   520
                 THEN ALLGOALS
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   521
                    (simp_tac (HOL_ss addsimps [real_minus_mult_eq2]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   522
                                               bin_simps@real_add_ac@real_mult_ac))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   523
  val numeral_simp_tac	= ALLGOALS 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   524
                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   525
  val simplify_meta_eq  = simplify_meta_eq
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   526
  end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   527
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   528
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   529
  
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   530
val combine_numerals = 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   531
    prep_simproc ("real_combine_numerals",
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   532
		  prep_pats ["(i::real) + j", "(i::real) - j"],
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   533
		  CombineNumerals.proc);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   534
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   535
end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   536
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   537
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   538
Addsimprocs Real_Numeral_Simprocs.cancel_numerals;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   539
Addsimprocs [Real_Numeral_Simprocs.combine_numerals];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   540
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   541
(*The Abel_Cancel simprocs are now obsolete*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   542
Delsimprocs [Real_Cancel.sum_conv, Real_Cancel.rel_conv];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   543
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   544
(*examples:
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   545
print_depth 22;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   546
set timing;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   547
set trace_simp;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   548
fun test s = (Goal s; by (Simp_tac 1)); 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   549
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   550
test "l + #2 + #2 + #2 + (l + #2) + (oo + #2) = (uu::real)";
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 "#2*u = (u::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   553
test "(i + j + #12 + (k::real)) - #15 = y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   554
test "(i + j + #12 + (k::real)) - #5 = y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   555
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   556
test "y - b < (b::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   557
test "y - (#3*b + c) < (b::real) - #2*c";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   558
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   559
test "(#2*x - (u*v) + y) - v*#3*u = (w::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   560
test "(#2*x*u*v + (u*v)*#4 + y) - v*u*#4 = (w::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   561
test "(#2*x*u*v + (u*v)*#4 + y) - v*u = (w::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   562
test "u*v - (x*u*v + (u*v)*#4 + y) = (w::real)";
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
test "(i + j + #12 + (k::real)) = u + #15 + y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   565
test "(i + j*#2 + #12 + (k::real)) = j + #5 + y";
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
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)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   568
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   569
test "a + -(b+c) + b = (d::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   570
test "a + -(b+c) - b = (d::real)";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   571
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   572
(*negative numerals*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   573
test "(i + j + #-2 + (k::real)) - (u + #5 + y) = zz";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   574
test "(i + j + #-3 + (k::real)) < u + #5 + y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   575
test "(i + j + #3 + (k::real)) < u + #-6 + y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   576
test "(i + j + #-12 + (k::real)) - #15 = y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   577
test "(i + j + #12 + (k::real)) - #-15 = y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   578
test "(i + j + #-12 + (k::real)) - #-15 = y";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   579
*)
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
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   582
(** Constant folding for real plus and times **)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   583
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   584
(*We do not need
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   585
    structure Real_Plus_Assoc = Assoc_Fold (Real_Plus_Assoc_Data);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   586
  because combine_numerals does the same thing*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   587
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   588
structure Real_Times_Assoc_Data : ASSOC_FOLD_DATA =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   589
struct
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   590
  val ss		= HOL_ss
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   591
  val eq_reflection	= eq_reflection
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   592
  val thy    = RealBin.thy
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   593
  val T	     = HOLogic.realT
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   594
  val plus   = Const ("op *", [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   595
  val add_ac = real_mult_ac
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   596
end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   597
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   598
structure Real_Times_Assoc = Assoc_Fold (Real_Times_Assoc_Data);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   599
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   600
Addsimprocs [Real_Times_Assoc.conv];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   601
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   602
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   603
(*** decision procedure for linear arithmetic ***)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   604
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   605
(*---------------------------------------------------------------------------*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   606
(* Linear arithmetic                                                         *)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   607
(*---------------------------------------------------------------------------*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   608
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   609
(*
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   610
Instantiation of the generic linear arithmetic package for real.
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   611
*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   612
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   613
(* Update parameters of arithmetic prover *)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   614
let
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   615
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   616
(* reduce contradictory <= to False *)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   617
val add_rules =  
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   618
    real_diff_def ::
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   619
    map (rename_numerals thy) 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   620
        [real_add_zero_left, real_add_zero_right, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   621
	 real_add_minus, real_add_minus_left, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   622
	 real_mult_0, real_mult_0_right, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   623
	 real_mult_1, real_mult_1_right, 
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   624
	 real_mult_minus_1, real_mult_minus_1_right];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   625
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   626
val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals]@
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   627
               Real_Numeral_Simprocs.cancel_numerals;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   628
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   629
val add_mono_thms =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   630
  map (fn s => prove_goal RealBin.thy s
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   631
                 (fn prems => [cut_facts_tac prems 1,
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   632
                      asm_simp_tac (simpset() addsimps [real_add_le_mono]) 1]))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   633
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::real)",
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   634
     "(i  = j) & (k <= l) ==> i + k <= j + (l::real)",
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   635
     "(i <= j) & (k  = l) ==> i + k <= j + (l::real)",
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   636
     "(i  = j) & (k  = l) ==> i + k  = j + (l::real)"
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   637
    ];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   638
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   639
in
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   640
LA_Data_Ref.add_mono_thms := !LA_Data_Ref.add_mono_thms @ add_mono_thms;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   641
(*We don't change LA_Data_Ref.lessD and LA_Data_Ref.discrete
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   642
 because the real ordering is dense!*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   643
LA_Data_Ref.ss_ref := !LA_Data_Ref.ss_ref addsimps add_rules
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   644
                      addsimprocs simprocs
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   645
                      addcongs [if_weak_cong]
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   646
end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   647
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   648
let
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   649
val real_arith_simproc_pats =
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   650
  map (fn s => Thm.read_cterm (Theory.sign_of RealDef.thy) (s, HOLogic.boolT))
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   651
      ["(m::real) < n","(m::real) <= n", "(m::real) = n"];
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   652
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   653
val fast_real_arith_simproc = mk_simproc
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   654
  "fast_real_arith" real_arith_simproc_pats Fast_Arith.lin_arith_prover;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   655
in
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   656
Addsimprocs [fast_real_arith_simproc]
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   657
end;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   658
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   659
(* Some test data [omitting examples thet assume the ordering to be discrete!]
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   660
Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   661
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   662
Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   663
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   664
Goal "!!a::real. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   665
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   666
Goal "!!a::real. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   667
by (arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   668
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   669
\     ==> a <= l";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   670
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   671
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   672
\     ==> a+a+a+a <= l+l+l+l";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   673
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   674
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   675
\     ==> a+a+a+a+a <= l+l+l+l+i";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   676
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   677
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   678
\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   679
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   680
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   681
\     ==> #6*a <= #5*l+i";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   682
by (fast_arith_tac 1);
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   683
*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   684
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   685
(*---------------------------------------------------------------------------*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   686
(* End of linear arithmetic                                                  *)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   687
(*---------------------------------------------------------------------------*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   688
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   689
(*useful??*)
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   690
Goal "(z = z + w) = (w = (#0::real))";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   691
by Auto_tac;
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   692
qed "real_add_left_cancel0";
202fdf72cf23 installing the cancel_numerals and combine_numerals simprocs
paulson
parents: 9043
diff changeset
   693
Addsimps [real_add_left_cancel0];