author | paulson |
Mon, 22 Dec 2003 12:50:01 +0100 | |
changeset 14308 | c0489217deb2 |
parent 14305 | f17ca9f6dc8c |
child 14311 | efda5615bb7d |
permissions | -rw-r--r-- |
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Real/real_arith.ML |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
3 |
Author: Tobias Nipkow, TU Muenchen |
14289 | 4 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
5 |
|
14289 | 6 |
Simprocs: Common factor cancellation & Rational coefficient handling |
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
7 |
*) |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
8 |
|
14289 | 9 |
|
10 |
(** Misc ML bindings **) |
|
11 |
||
14308 | 12 |
val inverse_less_iff_less = thm"inverse_less_iff_less"; |
14289 | 13 |
|
14 |
val pos_real_less_divide_eq = thm"pos_less_divide_eq"; |
|
15 |
val pos_real_divide_less_eq = thm"pos_divide_less_eq"; |
|
16 |
val pos_real_le_divide_eq = thm"pos_le_divide_eq"; |
|
17 |
val pos_real_divide_le_eq = thm"pos_divide_le_eq"; |
|
18 |
||
19 |
||
20 |
(****Common factor cancellation****) |
|
21 |
||
14293 | 22 |
(*To quote from Provers/Arith/cancel_numeral_factor.ML: |
23 |
||
24 |
This simproc Cancels common coefficients in balanced expressions: |
|
25 |
||
26 |
u*#m ~~ u'*#m' == #n*u ~~ #n'*u' |
|
27 |
||
28 |
where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /) |
|
29 |
and d = gcd(m,m') and n=m/d and n'=m'/d. |
|
30 |
*) |
|
31 |
||
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
32 |
val inverse_eq_divide = thm"Ring_and_Field.inverse_eq_divide"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
33 |
|
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
34 |
val mult_less_cancel_left = thm"Ring_and_Field.mult_less_cancel_left"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
35 |
val mult_le_cancel_left = thm"Ring_and_Field.mult_le_cancel_left"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
36 |
val mult_less_cancel_right = thm"Ring_and_Field.mult_less_cancel_right"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
37 |
val mult_le_cancel_right = thm"Ring_and_Field.mult_le_cancel_right"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
38 |
val mult_cancel_left = thm"Ring_and_Field.mult_cancel_left"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
39 |
val mult_cancel_right = thm"Ring_and_Field.mult_cancel_right"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
40 |
|
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
41 |
val mult_divide_cancel_left = thm"Ring_and_Field.mult_divide_cancel_left"; |
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
42 |
val mult_divide_cancel_eq_if = thm"Ring_and_Field.mult_divide_cancel_eq_if"; |
14289 | 43 |
|
14288 | 44 |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
45 |
local |
14289 | 46 |
open Real_Numeral_Simprocs |
47 |
in |
|
48 |
||
49 |
val rel_real_number_of = [eq_real_number_of, less_real_number_of, |
|
50 |
le_real_number_of_eq_not_less] |
|
51 |
||
52 |
structure CancelNumeralFactorCommon = |
|
53 |
struct |
|
54 |
val mk_coeff = mk_coeff |
|
55 |
val dest_coeff = dest_coeff 1 |
|
56 |
val trans_tac = trans_tac |
|
57 |
val norm_tac = |
|
58 |
ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps @ mult_1s)) |
|
59 |
THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@real_mult_minus_simps)) |
|
60 |
THEN ALLGOALS (simp_tac (HOL_ss addsimps real_mult_ac)) |
|
61 |
val numeral_simp_tac = |
|
62 |
ALLGOALS (simp_tac (HOL_ss addsimps rel_real_number_of@bin_simps)) |
|
63 |
val simplify_meta_eq = simplify_meta_eq |
|
64 |
end |
|
65 |
||
66 |
structure DivCancelNumeralFactor = CancelNumeralFactorFun |
|
67 |
(open CancelNumeralFactorCommon |
|
68 |
val prove_conv = Bin_Simprocs.prove_conv |
|
69 |
val mk_bal = HOLogic.mk_binop "HOL.divide" |
|
70 |
val dest_bal = HOLogic.dest_bin "HOL.divide" HOLogic.realT |
|
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
71 |
val cancel = mult_divide_cancel_left RS trans |
14289 | 72 |
val neg_exchanges = false |
73 |
) |
|
74 |
||
75 |
structure EqCancelNumeralFactor = CancelNumeralFactorFun |
|
76 |
(open CancelNumeralFactorCommon |
|
77 |
val prove_conv = Bin_Simprocs.prove_conv |
|
78 |
val mk_bal = HOLogic.mk_eq |
|
79 |
val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT |
|
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
80 |
val cancel = mult_cancel_left RS trans |
14289 | 81 |
val neg_exchanges = false |
82 |
) |
|
83 |
||
84 |
structure LessCancelNumeralFactor = CancelNumeralFactorFun |
|
85 |
(open CancelNumeralFactorCommon |
|
86 |
val prove_conv = Bin_Simprocs.prove_conv |
|
87 |
val mk_bal = HOLogic.mk_binrel "op <" |
|
88 |
val dest_bal = HOLogic.dest_bin "op <" HOLogic.realT |
|
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
89 |
val cancel = mult_less_cancel_left RS trans |
14289 | 90 |
val neg_exchanges = true |
91 |
) |
|
92 |
||
93 |
structure LeCancelNumeralFactor = CancelNumeralFactorFun |
|
94 |
(open CancelNumeralFactorCommon |
|
95 |
val prove_conv = Bin_Simprocs.prove_conv |
|
96 |
val mk_bal = HOLogic.mk_binrel "op <=" |
|
97 |
val dest_bal = HOLogic.dest_bin "op <=" HOLogic.realT |
|
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
98 |
val cancel = mult_le_cancel_left RS trans |
14289 | 99 |
val neg_exchanges = true |
100 |
) |
|
101 |
||
102 |
val real_cancel_numeral_factors_relations = |
|
103 |
map prep_simproc |
|
104 |
[("realeq_cancel_numeral_factor", |
|
105 |
["(l::real) * m = n", "(l::real) = m * n"], |
|
106 |
EqCancelNumeralFactor.proc), |
|
107 |
("realless_cancel_numeral_factor", |
|
108 |
["(l::real) * m < n", "(l::real) < m * n"], |
|
109 |
LessCancelNumeralFactor.proc), |
|
110 |
("realle_cancel_numeral_factor", |
|
111 |
["(l::real) * m <= n", "(l::real) <= m * n"], |
|
112 |
LeCancelNumeralFactor.proc)] |
|
113 |
||
114 |
val real_cancel_numeral_factors_divide = prep_simproc |
|
115 |
("realdiv_cancel_numeral_factor", |
|
116 |
["((l::real) * m) / n", "(l::real) / (m * n)", |
|
117 |
"((number_of v)::real) / (number_of w)"], |
|
118 |
DivCancelNumeralFactor.proc) |
|
119 |
||
120 |
val real_cancel_numeral_factors = |
|
121 |
real_cancel_numeral_factors_relations @ |
|
122 |
[real_cancel_numeral_factors_divide] |
|
123 |
||
124 |
end; |
|
125 |
||
126 |
Addsimprocs real_cancel_numeral_factors; |
|
127 |
||
128 |
||
129 |
(*examples: |
|
130 |
print_depth 22; |
|
131 |
set timing; |
|
132 |
set trace_simp; |
|
133 |
fun test s = (Goal s; by (Simp_tac 1)); |
|
134 |
||
135 |
test "0 <= (y::real) * -2"; |
|
136 |
test "9*x = 12 * (y::real)"; |
|
137 |
test "(9*x) / (12 * (y::real)) = z"; |
|
138 |
test "9*x < 12 * (y::real)"; |
|
139 |
test "9*x <= 12 * (y::real)"; |
|
140 |
||
141 |
test "-99*x = 132 * (y::real)"; |
|
142 |
test "(-99*x) / (132 * (y::real)) = z"; |
|
143 |
test "-99*x < 132 * (y::real)"; |
|
144 |
test "-99*x <= 132 * (y::real)"; |
|
145 |
||
146 |
test "999*x = -396 * (y::real)"; |
|
147 |
test "(999*x) / (-396 * (y::real)) = z"; |
|
148 |
test "999*x < -396 * (y::real)"; |
|
149 |
test "999*x <= -396 * (y::real)"; |
|
150 |
||
151 |
test "(- ((2::real) * x) <= 2 * y)"; |
|
152 |
test "-99*x = -81 * (y::real)"; |
|
153 |
test "(-99*x) / (-81 * (y::real)) = z"; |
|
154 |
test "-99*x <= -81 * (y::real)"; |
|
155 |
test "-99*x < -81 * (y::real)"; |
|
156 |
||
157 |
test "-2 * x = -1 * (y::real)"; |
|
158 |
test "-2 * x = -(y::real)"; |
|
159 |
test "(-2 * x) / (-1 * (y::real)) = z"; |
|
160 |
test "-2 * x < -(y::real)"; |
|
161 |
test "-2 * x <= -1 * (y::real)"; |
|
162 |
test "-x < -23 * (y::real)"; |
|
163 |
test "-x <= -23 * (y::real)"; |
|
164 |
*) |
|
165 |
||
166 |
||
167 |
(** Declarations for ExtractCommonTerm **) |
|
168 |
||
169 |
local |
|
170 |
open Real_Numeral_Simprocs |
|
171 |
in |
|
172 |
||
173 |
structure CancelFactorCommon = |
|
174 |
struct |
|
175 |
val mk_sum = long_mk_prod |
|
176 |
val dest_sum = dest_prod |
|
177 |
val mk_coeff = mk_coeff |
|
178 |
val dest_coeff = dest_coeff |
|
179 |
val find_first = find_first [] |
|
180 |
val trans_tac = trans_tac |
|
181 |
val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@real_mult_ac)) |
|
182 |
end; |
|
183 |
||
184 |
structure EqCancelFactor = ExtractCommonTermFun |
|
185 |
(open CancelFactorCommon |
|
186 |
val prove_conv = Bin_Simprocs.prove_conv |
|
187 |
val mk_bal = HOLogic.mk_eq |
|
188 |
val dest_bal = HOLogic.dest_bin "op =" HOLogic.realT |
|
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
189 |
val simplify_meta_eq = cancel_simplify_meta_eq mult_cancel_left |
14289 | 190 |
); |
191 |
||
192 |
||
193 |
structure DivideCancelFactor = ExtractCommonTermFun |
|
194 |
(open CancelFactorCommon |
|
195 |
val prove_conv = Bin_Simprocs.prove_conv |
|
196 |
val mk_bal = HOLogic.mk_binop "HOL.divide" |
|
197 |
val dest_bal = HOLogic.dest_bin "HOL.divide" HOLogic.realT |
|
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14293
diff
changeset
|
198 |
val simplify_meta_eq = cancel_simplify_meta_eq mult_divide_cancel_eq_if |
14289 | 199 |
); |
200 |
||
201 |
val real_cancel_factor = |
|
202 |
map prep_simproc |
|
203 |
[("real_eq_cancel_factor", ["(l::real) * m = n", "(l::real) = m * n"], EqCancelFactor.proc), |
|
204 |
("real_divide_cancel_factor", ["((l::real) * m) / n", "(l::real) / (m * n)"], |
|
205 |
DivideCancelFactor.proc)]; |
|
206 |
||
207 |
end; |
|
208 |
||
209 |
Addsimprocs real_cancel_factor; |
|
210 |
||
211 |
||
212 |
(*examples: |
|
213 |
print_depth 22; |
|
214 |
set timing; |
|
215 |
set trace_simp; |
|
216 |
fun test s = (Goal s; by (Asm_simp_tac 1)); |
|
217 |
||
218 |
test "x*k = k*(y::real)"; |
|
219 |
test "k = k*(y::real)"; |
|
220 |
test "a*(b*c) = (b::real)"; |
|
221 |
test "a*(b*c) = d*(b::real)*(x*a)"; |
|
222 |
||
223 |
||
224 |
test "(x*k) / (k*(y::real)) = (uu::real)"; |
|
225 |
test "(k) / (k*(y::real)) = (uu::real)"; |
|
226 |
test "(a*(b*c)) / ((b::real)) = (uu::real)"; |
|
227 |
test "(a*(b*c)) / (d*(b::real)*(x*a)) = (uu::real)"; |
|
228 |
||
229 |
(*FIXME: what do we do about this?*) |
|
230 |
test "a*(b*c)/(y*z) = d*(b::real)*(x*a)/z"; |
|
231 |
*) |
|
232 |
||
233 |
(****Augmentation of real linear arithmetic with |
|
234 |
rational coefficient handling****) |
|
235 |
||
236 |
val divide_1 = thm"divide_1"; |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
237 |
|
14288 | 238 |
val times_divide_eq_left = thm"times_divide_eq_left"; |
239 |
val times_divide_eq_right = thm"times_divide_eq_right"; |
|
240 |
||
14289 | 241 |
local |
242 |
||
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
243 |
(* reduce contradictory <= to False *) |
14288 | 244 |
val simps = [True_implies_equals, |
245 |
inst "w" "number_of ?v" real_add_mult_distrib2, |
|
14289 | 246 |
divide_1,times_divide_eq_right,times_divide_eq_left]; |
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
247 |
|
14293 | 248 |
val simprocs = [real_cancel_numeral_factors_divide, |
249 |
real_cancel_numeral_factors_divide]; |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
250 |
|
10693 | 251 |
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var; |
252 |
||
253 |
val real_mult_mono_thms = |
|
254 |
[(rotate_prems 1 real_mult_less_mono2, |
|
255 |
cvar(real_mult_less_mono2, hd(prems_of real_mult_less_mono2))), |
|
256 |
(real_mult_le_mono2, |
|
257 |
cvar(real_mult_le_mono2, hd(tl(prems_of real_mult_le_mono2))))] |
|
258 |
||
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
259 |
in |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
260 |
|
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
261 |
val real_arith_setup = |
10693 | 262 |
[Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} => |
10722 | 263 |
{add_mono_thms = add_mono_thms, |
10693 | 264 |
mult_mono_thms = mult_mono_thms @ real_mult_mono_thms, |
10722 | 265 |
inj_thms = inj_thms, |
266 |
lessD = lessD, |
|
267 |
simpset = simpset addsimps simps addsimprocs simprocs})]; |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
268 |
|
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
269 |
end; |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
270 |
|
10722 | 271 |
(* |
272 |
Procedure "assoc_fold" needed? |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
273 |
*) |