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