author | skalberg |
Sun, 13 Feb 2005 17:15:14 +0100 | |
changeset 15531 | 08c8dad8e399 |
parent 15186 | 1fb9a1fe8d0c |
child 15661 | 9ef583b08647 |
permissions | -rw-r--r-- |
14426 | 1 |
(* Title: HOL/Real/real_arith.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 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
11 |
val real_le_def = thm "real_le_def"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
12 |
val real_diff_def = thm "real_diff_def"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
13 |
val real_divide_def = thm "real_divide_def"; |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14356
diff
changeset
|
14 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
15 |
val realrel_in_real = thm"realrel_in_real"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
16 |
val real_add_commute = thm"real_add_commute"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
17 |
val real_add_assoc = thm"real_add_assoc"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
18 |
val real_add_zero_left = thm"real_add_zero_left"; |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14356
diff
changeset
|
19 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
20 |
val real_mult_commute = thm"real_mult_commute"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
21 |
val real_mult_assoc = thm"real_mult_assoc"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
22 |
val real_mult_1 = thm"real_mult_1"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
23 |
val real_mult_1_right = thm"real_mult_1_right"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
24 |
val preal_le_linear = thm"preal_le_linear"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
25 |
val real_mult_inverse_left = thm"real_mult_inverse_left"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
26 |
val real_not_refl2 = thm"real_not_refl2"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
27 |
val real_of_preal_add = thm"real_of_preal_add"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
28 |
val real_of_preal_mult = thm"real_of_preal_mult"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
29 |
val real_of_preal_trichotomy = thm"real_of_preal_trichotomy"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
30 |
val real_of_preal_minus_less_zero = thm"real_of_preal_minus_less_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
31 |
val real_of_preal_not_minus_gt_zero = thm"real_of_preal_not_minus_gt_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
32 |
val real_of_preal_zero_less = thm"real_of_preal_zero_less"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
33 |
val real_le_imp_less_or_eq = thm"real_le_imp_less_or_eq"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
34 |
val real_le_refl = thm"real_le_refl"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
35 |
val real_le_linear = thm"real_le_linear"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
36 |
val real_le_trans = thm"real_le_trans"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
37 |
val real_less_le = thm"real_less_le"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
38 |
val real_less_sum_gt_zero = thm"real_less_sum_gt_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
39 |
val real_gt_zero_preal_Ex = thm "real_gt_zero_preal_Ex"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
40 |
val real_gt_preal_preal_Ex = thm "real_gt_preal_preal_Ex"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
41 |
val real_ge_preal_preal_Ex = thm "real_ge_preal_preal_Ex"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
42 |
val real_less_all_preal = thm "real_less_all_preal"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
43 |
val real_less_all_real2 = thm "real_less_all_real2"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
44 |
val real_of_preal_le_iff = thm "real_of_preal_le_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
45 |
val real_mult_order = thm "real_mult_order"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
46 |
val real_add_less_le_mono = thm "real_add_less_le_mono"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
47 |
val real_add_le_less_mono = thm "real_add_le_less_mono"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
48 |
val real_add_order = thm "real_add_order"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
49 |
val real_le_add_order = thm "real_le_add_order"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
50 |
val real_le_square = thm "real_le_square"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
51 |
val real_mult_less_mono2 = thm "real_mult_less_mono2"; |
14289 | 52 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
53 |
val real_mult_less_iff1 = thm "real_mult_less_iff1"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
54 |
val real_mult_le_cancel_iff1 = thm "real_mult_le_cancel_iff1"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
55 |
val real_mult_le_cancel_iff2 = thm "real_mult_le_cancel_iff2"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
56 |
val real_mult_less_mono = thm "real_mult_less_mono"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
57 |
val real_mult_less_mono' = thm "real_mult_less_mono'"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
58 |
val real_sum_squares_cancel = thm "real_sum_squares_cancel"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
59 |
val real_sum_squares_cancel2 = thm "real_sum_squares_cancel2"; |
14289 | 60 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
61 |
val real_mult_left_cancel = thm"real_mult_left_cancel"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
62 |
val real_mult_right_cancel = thm"real_mult_right_cancel"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
63 |
val real_inverse_unique = thm "real_inverse_unique"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
64 |
val real_inverse_gt_one = thm "real_inverse_gt_one"; |
14289 | 65 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
66 |
val real_of_int_zero = thm"real_of_int_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
67 |
val real_of_one = thm"real_of_one"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
68 |
val real_of_int_add = thm"real_of_int_add"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
69 |
val real_of_int_minus = thm"real_of_int_minus"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
70 |
val real_of_int_diff = thm"real_of_int_diff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
71 |
val real_of_int_mult = thm"real_of_int_mult"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
72 |
val real_of_int_real_of_nat = thm"real_of_int_real_of_nat"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
73 |
val real_of_int_inject = thm"real_of_int_inject"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
74 |
val real_of_int_less_iff = thm"real_of_int_less_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
75 |
val real_of_int_le_iff = thm"real_of_int_le_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
76 |
val real_of_nat_zero = thm "real_of_nat_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
77 |
val real_of_nat_one = thm "real_of_nat_one"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
78 |
val real_of_nat_add = thm "real_of_nat_add"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
79 |
val real_of_nat_Suc = thm "real_of_nat_Suc"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
80 |
val real_of_nat_less_iff = thm "real_of_nat_less_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
81 |
val real_of_nat_le_iff = thm "real_of_nat_le_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
82 |
val real_of_nat_ge_zero = thm "real_of_nat_ge_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
83 |
val real_of_nat_Suc_gt_zero = thm "real_of_nat_Suc_gt_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
84 |
val real_of_nat_mult = thm "real_of_nat_mult"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
85 |
val real_of_nat_inject = thm "real_of_nat_inject"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
86 |
val real_of_nat_diff = thm "real_of_nat_diff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
87 |
val real_of_nat_zero_iff = thm "real_of_nat_zero_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
88 |
val real_of_nat_gt_zero_cancel_iff = thm "real_of_nat_gt_zero_cancel_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
89 |
val real_of_nat_le_zero_cancel_iff = thm "real_of_nat_le_zero_cancel_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
90 |
val not_real_of_nat_less_zero = thm "not_real_of_nat_less_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
91 |
val real_of_nat_ge_zero_cancel_iff = thm "real_of_nat_ge_zero_cancel_iff"; |
14390 | 92 |
val real_number_of = thm"real_number_of"; |
93 |
val real_of_nat_number_of = thm"real_of_nat_number_of"; |
|
14426 | 94 |
val real_of_int_of_nat_eq = thm"real_of_int_of_nat_eq"; |
14352 | 95 |
|
96 |
||
97 |
(****Instantiation of the generic linear arithmetic package****) |
|
98 |
||
14289 | 99 |
local |
100 |
||
14369 | 101 |
val simps = [real_of_nat_zero, real_of_nat_Suc, real_of_nat_add, |
102 |
real_of_nat_mult, real_of_int_zero, real_of_one, real_of_int_add RS sym, |
|
103 |
real_of_int_minus RS sym, real_of_int_diff RS sym, |
|
14426 | 104 |
real_of_int_mult RS sym, real_of_int_of_nat_eq, |
14390 | 105 |
real_of_nat_number_of, real_number_of]; |
14355
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents:
14352
diff
changeset
|
106 |
|
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents:
14352
diff
changeset
|
107 |
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
|
108 |
real_of_int_inject RS iffD2]; |
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents:
14352
diff
changeset
|
109 |
|
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents:
14352
diff
changeset
|
110 |
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
|
111 |
real_of_nat_inject RS iffD2]; |
14352 | 112 |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
113 |
in |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
114 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
115 |
val fast_real_arith_simproc = |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
116 |
Simplifier.simproc (Theory.sign_of (the_context ())) |
14352 | 117 |
"fast_real_arith" ["(m::real) < n","(m::real) <= n", "(m::real) = n"] |
118 |
Fast_Arith.lin_arith_prover; |
|
119 |
||
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
120 |
val real_arith_setup = |
10693 | 121 |
[Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} => |
14368
2763da611ad9
converted Real/Lubs to Isar script. Converting arithmetic setup
paulson
parents:
14365
diff
changeset
|
122 |
{add_mono_thms = add_mono_thms, |
15184 | 123 |
mult_mono_thms = mult_mono_thms, |
14355
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents:
14352
diff
changeset
|
124 |
inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms, |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14356
diff
changeset
|
125 |
lessD = lessD, (*Can't change LA_Data_Ref.lessD: the reals are dense!*) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14369
diff
changeset
|
126 |
simpset = simpset addsimps simps}), |
14355
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
nipkow
parents:
14352
diff
changeset
|
127 |
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
|
128 |
arith_inj_const ("RealDef.real", HOLogic.intT --> HOLogic.realT), |
14352 | 129 |
Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]]; |
130 |
||
131 |
(* some thms for injection nat => real: |
|
132 |
real_of_nat_zero |
|
133 |
real_of_nat_add |
|
134 |
*) |
|
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
135 |
|
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
136 |
end; |
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
137 |
|
14352 | 138 |
|
15531 | 139 |
(* SOME test data [omitting examples that assume the ordering to be discrete!] |
14352 | 140 |
Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d"; |
141 |
by (fast_arith_tac 1); |
|
142 |
qed ""; |
|
143 |
||
144 |
Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c"; |
|
145 |
by (fast_arith_tac 1); |
|
146 |
qed ""; |
|
147 |
||
148 |
Goal "!!a::real. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j"; |
|
149 |
by (fast_arith_tac 1); |
|
150 |
qed ""; |
|
151 |
||
152 |
Goal "!!a::real. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k"; |
|
153 |
by (arith_tac 1); |
|
154 |
qed ""; |
|
155 |
||
156 |
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ |
|
157 |
\ ==> a <= l"; |
|
158 |
by (fast_arith_tac 1); |
|
159 |
qed ""; |
|
160 |
||
161 |
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ |
|
162 |
\ ==> a+a+a+a <= l+l+l+l"; |
|
163 |
by (fast_arith_tac 1); |
|
164 |
qed ""; |
|
165 |
||
166 |
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ |
|
167 |
\ ==> a+a+a+a+a <= l+l+l+l+i"; |
|
168 |
by (fast_arith_tac 1); |
|
169 |
qed ""; |
|
170 |
||
171 |
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ |
|
172 |
\ ==> a+a+a+a+a+a <= l+l+l+l+i+l"; |
|
173 |
by (fast_arith_tac 1); |
|
174 |
qed ""; |
|
175 |
||
176 |
Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ |
|
177 |
\ ==> 6*a <= 5*l+i"; |
|
178 |
by (fast_arith_tac 1); |
|
179 |
qed ""; |
|
180 |
||
181 |
Goal "a<=b ==> a < b+(1::real)"; |
|
182 |
by (fast_arith_tac 1); |
|
183 |
qed ""; |
|
184 |
||
185 |
Goal "a<=b ==> a-(3::real) < b"; |
|
186 |
by (fast_arith_tac 1); |
|
187 |
qed ""; |
|
188 |
||
189 |
Goal "a<=b ==> a-(1::real) < b"; |
|
190 |
by (fast_arith_tac 1); |
|
191 |
qed ""; |
|
192 |
||
9436
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff
changeset
|
193 |
*) |
14352 | 194 |
|
195 |
||
196 |