author | wenzelm |
Wed, 05 Dec 2001 03:13:57 +0100 | |
changeset 12378 | 86c58273f8c0 |
parent 12018 | ec054019c910 |
child 14268 | 5cf13e80be0e |
permissions | -rw-r--r-- |
10751 | 1 |
(* Title: Real/Hyperreal/HyperOrd.ML |
2 |
Author: Jacques D. Fleuriot |
|
3 |
Copyright: 1998 University of Cambridge |
|
4 |
2000 University of Edinburgh |
|
5 |
Description: Type "hypreal" is a linear order and also |
|
6 |
satisfies plus_ac0: + is an AC-operator with zero |
|
7 |
*) |
|
8 |
||
9 |
(**** The simproc abel_cancel ****) |
|
10 |
||
11 |
(*** Two lemmas needed for the simprocs ***) |
|
12 |
||
13 |
(*Deletion of other terms in the formula, seeking the -x at the front of z*) |
|
14 |
Goal "((x::hypreal) + (y + z) = y + u) = ((x + z) = u)"; |
|
15 |
by (stac hypreal_add_left_commute 1); |
|
16 |
by (rtac hypreal_add_left_cancel 1); |
|
17 |
qed "hypreal_add_cancel_21"; |
|
18 |
||
19 |
(*A further rule to deal with the case that |
|
20 |
everything gets cancelled on the right.*) |
|
21 |
Goal "((x::hypreal) + (y + z) = y) = (x = -z)"; |
|
22 |
by (stac hypreal_add_left_commute 1); |
|
23 |
by (res_inst_tac [("t", "y")] (hypreal_add_zero_right RS subst) 1 |
|
24 |
THEN stac hypreal_add_left_cancel 1); |
|
25 |
by (simp_tac (simpset() addsimps [hypreal_eq_diff_eq RS sym]) 1); |
|
26 |
qed "hypreal_add_cancel_end"; |
|
27 |
||
28 |
||
29 |
structure Hyperreal_Cancel_Data = |
|
30 |
struct |
|
31 |
val ss = HOL_ss |
|
32 |
val eq_reflection = eq_reflection |
|
33 |
||
34 |
val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) |
|
35 |
val T = Type("HyperDef.hypreal",[]) |
|
36 |
val zero = Const ("0", T) |
|
37 |
val restrict_to_left = restrict_to_left |
|
38 |
val add_cancel_21 = hypreal_add_cancel_21 |
|
39 |
val add_cancel_end = hypreal_add_cancel_end |
|
40 |
val add_left_cancel = hypreal_add_left_cancel |
|
41 |
val add_assoc = hypreal_add_assoc |
|
42 |
val add_commute = hypreal_add_commute |
|
43 |
val add_left_commute = hypreal_add_left_commute |
|
44 |
val add_0 = hypreal_add_zero_left |
|
45 |
val add_0_right = hypreal_add_zero_right |
|
46 |
||
47 |
val eq_diff_eq = hypreal_eq_diff_eq |
|
48 |
val eqI_rules = [hypreal_less_eqI, hypreal_eq_eqI, hypreal_le_eqI] |
|
49 |
fun dest_eqI th = |
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
10919
diff
changeset
|
50 |
#1 (HOLogic.dest_bin "op =" HOLogic.boolT |
10751 | 51 |
(HOLogic.dest_Trueprop (concl_of th))) |
52 |
||
53 |
val diff_def = hypreal_diff_def |
|
54 |
val minus_add_distrib = hypreal_minus_add_distrib |
|
55 |
val minus_minus = hypreal_minus_minus |
|
56 |
val minus_0 = hypreal_minus_zero |
|
57 |
val add_inverses = [hypreal_add_minus, hypreal_add_minus_left] |
|
58 |
val cancel_simps = [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA] |
|
59 |
end; |
|
60 |
||
61 |
structure Hyperreal_Cancel = Abel_Cancel (Hyperreal_Cancel_Data); |
|
62 |
||
63 |
Addsimprocs [Hyperreal_Cancel.sum_conv, Hyperreal_Cancel.rel_conv]; |
|
64 |
||
65 |
Goal "- (z - y) = y - (z::hypreal)"; |
|
66 |
by (Simp_tac 1); |
|
67 |
qed "hypreal_minus_diff_eq"; |
|
68 |
Addsimps [hypreal_minus_diff_eq]; |
|
69 |
||
70 |
||
71 |
Goal "((x::hypreal) < y) = (-y < -x)"; |
|
72 |
by (stac hypreal_less_minus_iff 1); |
|
73 |
by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1); |
|
74 |
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
|
75 |
qed "hypreal_less_swap_iff"; |
|
76 |
||
77 |
Goalw [hypreal_zero_def] |
|
78 |
"((0::hypreal) < x) = (-x < x)"; |
|
79 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
80 |
by (auto_tac (claset(), simpset() addsimps [hypreal_less, hypreal_minus])); |
|
81 |
by (ALLGOALS(Ultra_tac)); |
|
82 |
qed "hypreal_gt_zero_iff"; |
|
83 |
||
84 |
Goal "(A::hypreal) < B ==> A + C < B + C"; |
|
85 |
by (res_inst_tac [("z","A")] eq_Abs_hypreal 1); |
|
86 |
by (res_inst_tac [("z","B")] eq_Abs_hypreal 1); |
|
87 |
by (res_inst_tac [("z","C")] eq_Abs_hypreal 1); |
|
88 |
by (auto_tac (claset() addSIs [exI], |
|
89 |
simpset() addsimps [hypreal_less_def,hypreal_add])); |
|
90 |
by (Ultra_tac 1); |
|
91 |
qed "hypreal_add_less_mono1"; |
|
92 |
||
93 |
Goal "!!(A::hypreal). A < B ==> C + A < C + B"; |
|
94 |
by (auto_tac (claset() addIs [hypreal_add_less_mono1], |
|
95 |
simpset() addsimps [hypreal_add_commute])); |
|
96 |
qed "hypreal_add_less_mono2"; |
|
97 |
||
98 |
Goal "(x < (0::hypreal)) = (x < -x)"; |
|
99 |
by (rtac (hypreal_minus_zero_less_iff RS subst) 1); |
|
100 |
by (stac hypreal_gt_zero_iff 1); |
|
101 |
by (Full_simp_tac 1); |
|
102 |
qed "hypreal_lt_zero_iff"; |
|
103 |
||
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
104 |
Goalw [hypreal_zero_def] "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y"; |
10751 | 105 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
106 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
107 |
by (auto_tac (claset(), |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
108 |
simpset() addsimps [hypreal_less_def,hypreal_add])); |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
109 |
by (auto_tac (claset() addSIs [exI], |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
110 |
simpset() addsimps [hypreal_less_def,hypreal_add])); |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
111 |
by (ultra_tac (claset() addIs [real_add_order], simpset()) 1); |
10751 | 112 |
qed "hypreal_add_order"; |
113 |
||
114 |
Goal "[| 0 < x; 0 <= y |] ==> (0::hypreal) < x + y"; |
|
115 |
by (auto_tac (claset() addDs [sym, order_le_imp_less_or_eq] |
|
116 |
addIs [hypreal_add_order], |
|
117 |
simpset())); |
|
118 |
qed "hypreal_add_order_le"; |
|
119 |
||
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
120 |
Goalw [hypreal_zero_def] "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y"; |
10751 | 121 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
122 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
123 |
by (auto_tac (claset() addSIs [exI], |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
124 |
simpset() addsimps [hypreal_less_def,hypreal_mult])); |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
125 |
by (ultra_tac (claset() addIs [real_mult_order], simpset()) 1); |
10751 | 126 |
qed "hypreal_mult_order"; |
127 |
||
128 |
Goal "[| x < 0; y < 0 |] ==> (0::hypreal) < x * y"; |
|
129 |
by (REPEAT(dtac (hypreal_minus_zero_less_iff RS iffD2) 1)); |
|
130 |
by (dtac hypreal_mult_order 1 THEN assume_tac 1); |
|
131 |
by (Asm_full_simp_tac 1); |
|
132 |
qed "hypreal_mult_less_zero1"; |
|
133 |
||
134 |
Goal "[| 0 < x; y < 0 |] ==> x*y < (0::hypreal)"; |
|
135 |
by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1); |
|
136 |
by (dtac hypreal_mult_order 1 THEN assume_tac 1); |
|
137 |
by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1); |
|
138 |
by (Asm_full_simp_tac 1); |
|
139 |
qed "hypreal_mult_less_zero"; |
|
140 |
||
11713
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents:
11701
diff
changeset
|
141 |
Goalw [hypreal_one_def,hypreal_zero_def,hypreal_less_def] "0 < (1::hypreal)"; |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
142 |
by (res_inst_tac [("x","%n. 0")] exI 1); |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
143 |
by (res_inst_tac [("x","%n. 1")] exI 1); |
10751 | 144 |
by (auto_tac (claset(), |
145 |
simpset() addsimps [real_zero_less_one, FreeUltrafilterNat_Nat_set])); |
|
146 |
qed "hypreal_zero_less_one"; |
|
147 |
||
148 |
Goal "[| 0 <= x; 0 <= y |] ==> (0::hypreal) <= x + y"; |
|
149 |
by (REPEAT(dtac order_le_imp_less_or_eq 1)); |
|
150 |
by (auto_tac (claset() addIs [hypreal_add_order, order_less_imp_le], |
|
151 |
simpset())); |
|
152 |
qed "hypreal_le_add_order"; |
|
153 |
||
154 |
(*** Monotonicity results ***) |
|
155 |
||
156 |
Goal "(v+z < w+z) = (v < (w::hypreal))"; |
|
157 |
by (Simp_tac 1); |
|
158 |
qed "hypreal_add_right_cancel_less"; |
|
159 |
||
160 |
Goal "(z+v < z+w) = (v < (w::hypreal))"; |
|
161 |
by (Simp_tac 1); |
|
162 |
qed "hypreal_add_left_cancel_less"; |
|
163 |
||
164 |
Addsimps [hypreal_add_right_cancel_less, |
|
165 |
hypreal_add_left_cancel_less]; |
|
166 |
||
167 |
Goal "(v+z <= w+z) = (v <= (w::hypreal))"; |
|
168 |
by (Simp_tac 1); |
|
169 |
qed "hypreal_add_right_cancel_le"; |
|
170 |
||
171 |
Goal "(z+v <= z+w) = (v <= (w::hypreal))"; |
|
172 |
by (Simp_tac 1); |
|
173 |
qed "hypreal_add_left_cancel_le"; |
|
174 |
||
175 |
Addsimps [hypreal_add_right_cancel_le, hypreal_add_left_cancel_le]; |
|
176 |
||
177 |
Goal "[| (z1::hypreal) < y1; z2 < y2 |] ==> z1 + z2 < y1 + y2"; |
|
178 |
by (dtac (hypreal_less_minus_iff RS iffD1) 1); |
|
179 |
by (dtac (hypreal_less_minus_iff RS iffD1) 1); |
|
180 |
by (dtac hypreal_add_order 1 THEN assume_tac 1); |
|
181 |
by (thin_tac "0 < y2 + - z2" 1); |
|
182 |
by (dres_inst_tac [("C","z1 + z2")] hypreal_add_less_mono1 1); |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
183 |
by (auto_tac (claset(), |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
184 |
simpset() addsimps [hypreal_minus_add_distrib RS sym] @ hypreal_add_ac |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
185 |
delsimps [hypreal_minus_add_distrib])); |
10751 | 186 |
qed "hypreal_add_less_mono"; |
187 |
||
188 |
Goal "(q1::hypreal) <= q2 ==> x + q1 <= x + q2"; |
|
189 |
by (dtac order_le_imp_less_or_eq 1); |
|
190 |
by (Step_tac 1); |
|
191 |
by (auto_tac (claset() addSIs [order_less_imp_le,hypreal_add_less_mono1], |
|
192 |
simpset() addsimps [hypreal_add_commute])); |
|
193 |
qed "hypreal_add_left_le_mono1"; |
|
194 |
||
195 |
Goal "(q1::hypreal) <= q2 ==> q1 + x <= q2 + x"; |
|
196 |
by (auto_tac (claset() addDs [hypreal_add_left_le_mono1], |
|
197 |
simpset() addsimps [hypreal_add_commute])); |
|
198 |
qed "hypreal_add_le_mono1"; |
|
199 |
||
200 |
Goal "[|(i::hypreal)<=j; k<=l |] ==> i + k <= j + l"; |
|
201 |
by (etac (hypreal_add_le_mono1 RS order_trans) 1); |
|
202 |
by (Simp_tac 1); |
|
203 |
qed "hypreal_add_le_mono"; |
|
204 |
||
205 |
Goal "[|(i::hypreal)<j; k<=l |] ==> i + k < j + l"; |
|
206 |
by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] |
|
207 |
addIs [hypreal_add_less_mono1,hypreal_add_less_mono], |
|
208 |
simpset())); |
|
209 |
qed "hypreal_add_less_le_mono"; |
|
210 |
||
211 |
Goal "[|(i::hypreal)<=j; k<l |] ==> i + k < j + l"; |
|
212 |
by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] |
|
213 |
addIs [hypreal_add_less_mono2,hypreal_add_less_mono], |
|
214 |
simpset())); |
|
215 |
qed "hypreal_add_le_less_mono"; |
|
216 |
||
217 |
Goal "(A::hypreal) + C < B + C ==> A < B"; |
|
218 |
by (Full_simp_tac 1); |
|
219 |
qed "hypreal_less_add_right_cancel"; |
|
220 |
||
221 |
Goal "(C::hypreal) + A < C + B ==> A < B"; |
|
222 |
by (Full_simp_tac 1); |
|
223 |
qed "hypreal_less_add_left_cancel"; |
|
224 |
||
225 |
Goal "[|r < x; (0::hypreal) <= y|] ==> r < x + y"; |
|
226 |
by (auto_tac (claset() addDs [hypreal_add_less_le_mono], |
|
227 |
simpset())); |
|
228 |
qed "hypreal_add_zero_less_le_mono"; |
|
229 |
||
230 |
Goal "!!(A::hypreal). A + C <= B + C ==> A <= B"; |
|
231 |
by (dres_inst_tac [("x","-C")] hypreal_add_le_mono1 1); |
|
232 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1); |
|
233 |
qed "hypreal_le_add_right_cancel"; |
|
234 |
||
235 |
Goal "!!(A::hypreal). C + A <= C + B ==> A <= B"; |
|
236 |
by (dres_inst_tac [("x","-C")] hypreal_add_left_le_mono1 1); |
|
237 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
|
238 |
qed "hypreal_le_add_left_cancel"; |
|
239 |
||
240 |
Goal "(0::hypreal) <= x*x"; |
|
241 |
by (res_inst_tac [("x","0"),("y","x")] hypreal_linear_less2 1); |
|
242 |
by (auto_tac (claset() addIs [hypreal_mult_order, |
|
243 |
hypreal_mult_less_zero1,order_less_imp_le], |
|
244 |
simpset())); |
|
245 |
qed "hypreal_le_square"; |
|
246 |
Addsimps [hypreal_le_square]; |
|
247 |
||
248 |
Goalw [hypreal_le_def] "- (x*x) <= (0::hypreal)"; |
|
249 |
by (auto_tac (claset() addSDs [hypreal_le_square RS order_le_less_trans], |
|
250 |
simpset() addsimps [hypreal_minus_zero_less_iff])); |
|
251 |
qed "hypreal_less_minus_square"; |
|
252 |
Addsimps [hypreal_less_minus_square]; |
|
253 |
||
254 |
Goal "(0*x<r)=((0::hypreal)<r)"; |
|
255 |
by (Simp_tac 1); |
|
256 |
qed "hypreal_mult_0_less"; |
|
257 |
||
258 |
Goal "[| (0::hypreal) < z; x < y |] ==> x*z < y*z"; |
|
259 |
by (rotate_tac 1 1); |
|
260 |
by (dtac (hypreal_less_minus_iff RS iffD1) 1); |
|
261 |
by (rtac (hypreal_less_minus_iff RS iffD2) 1); |
|
262 |
by (dtac hypreal_mult_order 1 THEN assume_tac 1); |
|
263 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_mult_distrib2, |
|
264 |
hypreal_mult_commute ]) 1); |
|
265 |
qed "hypreal_mult_less_mono1"; |
|
266 |
||
267 |
Goal "[| (0::hypreal)<z; x<y |] ==> z*x<z*y"; |
|
268 |
by (asm_simp_tac (simpset() addsimps [hypreal_mult_commute,hypreal_mult_less_mono1]) 1); |
|
269 |
qed "hypreal_mult_less_mono2"; |
|
270 |
||
271 |
Goal "[| (0::hypreal)<=z; x<y |] ==> x*z<=y*z"; |
|
272 |
by (EVERY1 [rtac hypreal_less_or_eq_imp_le, dtac order_le_imp_less_or_eq]); |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
273 |
by (auto_tac (claset() addIs [hypreal_mult_less_mono1], simpset())); |
10751 | 274 |
qed "hypreal_mult_le_less_mono1"; |
275 |
||
276 |
Goal "[| (0::hypreal)<=z; x<y |] ==> z*x<=z*y"; |
|
277 |
by (asm_simp_tac (simpset() addsimps [hypreal_mult_commute, |
|
278 |
hypreal_mult_le_less_mono1]) 1); |
|
279 |
qed "hypreal_mult_le_less_mono2"; |
|
280 |
||
281 |
val prem1::prem2::prem3::rest = goal thy |
|
282 |
"[| (0::hypreal)<y; x<r; y*r<t*s |] ==> y*x<t*s"; |
|
283 |
by (rtac ([[prem1,prem2] MRS hypreal_mult_less_mono2, prem3] |
|
284 |
MRS order_less_trans) 1); |
|
285 |
qed "hypreal_mult_less_trans"; |
|
286 |
||
287 |
Goal "[| 0<=y; x<r; y*r<t*s; (0::hypreal)<t*s|] ==> y*x<t*s"; |
|
288 |
by (dtac order_le_imp_less_or_eq 1); |
|
289 |
by (fast_tac (HOL_cs addEs [hypreal_mult_0_less RS iffD2, |
|
290 |
hypreal_mult_less_trans]) 1); |
|
291 |
qed "hypreal_mult_le_less_trans"; |
|
292 |
||
293 |
Goal "[| 0 <= y; x <= r; y*r < t*s; (0::hypreal) < t*s|] ==> y*x < t*s"; |
|
294 |
by (dres_inst_tac [("x","x")] order_le_imp_less_or_eq 1); |
|
295 |
by (fast_tac (claset() addIs [hypreal_mult_le_less_trans]) 1); |
|
296 |
qed "hypreal_mult_le_le_trans"; |
|
297 |
||
298 |
Goal "[| u<v; x<y; (0::hypreal) < v; 0 < x |] ==> u*x < v* y"; |
|
299 |
by (etac (hypreal_mult_less_mono1 RS order_less_trans) 1); |
|
300 |
by (assume_tac 1); |
|
301 |
by (etac hypreal_mult_less_mono2 1); |
|
302 |
by (assume_tac 1); |
|
303 |
qed "hypreal_mult_less_mono"; |
|
304 |
||
305 |
(*UNUSED at present but possibly more useful than hypreal_mult_less_mono*) |
|
306 |
Goal "[| x < y; r1 < r2; (0::hypreal) <= r1; 0 <= x|] ==> r1 * x < r2 * y"; |
|
307 |
by (subgoal_tac "0<r2" 1); |
|
308 |
by (blast_tac (claset() addIs [order_le_less_trans]) 2); |
|
309 |
by (case_tac "x=0" 1); |
|
310 |
by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] |
|
311 |
addIs [hypreal_mult_less_mono, hypreal_mult_order], |
|
312 |
simpset())); |
|
313 |
qed "hypreal_mult_less_mono'"; |
|
314 |
||
315 |
Goal "0 < x ==> 0 < inverse (x::hypreal)"; |
|
316 |
by (EVERY1[rtac ccontr, dtac hypreal_leI]); |
|
317 |
by (forward_tac [hypreal_minus_zero_less_iff2 RS iffD2] 1); |
|
318 |
by (forward_tac [hypreal_not_refl2 RS not_sym] 1); |
|
319 |
by (dtac (hypreal_not_refl2 RS not_sym RS hypreal_inverse_not_zero) 1); |
|
320 |
by (EVERY1[dtac order_le_imp_less_or_eq, Step_tac]); |
|
321 |
by (dtac hypreal_mult_less_zero1 1 THEN assume_tac 1); |
|
322 |
by (auto_tac (claset() addIs [hypreal_zero_less_one RS hypreal_less_asym], |
|
323 |
simpset() addsimps [hypreal_minus_zero_less_iff])); |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
324 |
qed "hypreal_inverse_gt_0"; |
10751 | 325 |
|
326 |
Goal "x < 0 ==> inverse (x::hypreal) < 0"; |
|
327 |
by (ftac hypreal_not_refl2 1); |
|
328 |
by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1); |
|
329 |
by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1); |
|
330 |
by (stac (hypreal_minus_inverse RS sym) 1); |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
331 |
by (auto_tac (claset() addIs [hypreal_inverse_gt_0], simpset())); |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
332 |
qed "hypreal_inverse_less_0"; |
10751 | 333 |
|
334 |
Goal "(x::hypreal)*x <= x*x + y*y"; |
|
335 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
336 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
337 |
by (auto_tac (claset(), |
|
338 |
simpset() addsimps [hypreal_mult,hypreal_add,hypreal_le])); |
|
339 |
qed "hypreal_self_le_add_pos"; |
|
340 |
Addsimps [hypreal_self_le_add_pos]; |
|
341 |
||
342 |
(*lcp: new lemma unfortunately needed...*) |
|
343 |
Goal "-(x*x) <= (y*y::real)"; |
|
344 |
by (rtac order_trans 1); |
|
345 |
by (rtac real_le_square 2); |
|
346 |
by Auto_tac; |
|
347 |
qed "minus_square_le_square"; |
|
348 |
||
349 |
Goal "(x::hypreal)*x <= x*x + y*y + z*z"; |
|
350 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
351 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
352 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
|
353 |
by (auto_tac (claset(), |
|
354 |
simpset() addsimps [hypreal_mult, hypreal_add, hypreal_le, |
|
355 |
minus_square_le_square])); |
|
356 |
qed "hypreal_self_le_add_pos2"; |
|
357 |
Addsimps [hypreal_self_le_add_pos2]; |
|
358 |
||
359 |
||
360 |
(*---------------------------------------------------------------------------- |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
361 |
Existence of infinite hyperreal number |
10751 | 362 |
----------------------------------------------------------------------------*) |
363 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
364 |
Goalw [omega_def] "Rep_hypreal(omega) : hypreal"; |
10751 | 365 |
by (rtac Rep_hypreal 1); |
366 |
qed "Rep_hypreal_omega"; |
|
367 |
||
368 |
(* existence of infinite number not corresponding to any real number *) |
|
369 |
(* use assumption that member FreeUltrafilterNat is not finite *) |
|
370 |
(* a few lemmas first *) |
|
371 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
372 |
Goal "{n::nat. x = real n} = {} | \ |
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
373 |
\ (EX y. {n::nat. x = real n} = {y})"; |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
374 |
by (auto_tac (claset() addDs [inj_real_of_nat RS injD], simpset())); |
10751 | 375 |
qed "lemma_omega_empty_singleton_disj"; |
376 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
377 |
Goal "finite {n::nat. x = real n}"; |
10751 | 378 |
by (cut_inst_tac [("x","x")] lemma_omega_empty_singleton_disj 1); |
379 |
by Auto_tac; |
|
380 |
qed "lemma_finite_omega_set"; |
|
381 |
||
382 |
Goalw [omega_def,hypreal_of_real_def] |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
383 |
"~ (EX x. hypreal_of_real x = omega)"; |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
384 |
by (auto_tac (claset(), |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
385 |
simpset() addsimps [real_of_nat_Suc, real_diff_eq_eq RS sym, |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
386 |
lemma_finite_omega_set RS FreeUltrafilterNat_finite])); |
10751 | 387 |
qed "not_ex_hypreal_of_real_eq_omega"; |
388 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
389 |
Goal "hypreal_of_real x ~= omega"; |
10751 | 390 |
by (cut_facts_tac [not_ex_hypreal_of_real_eq_omega] 1); |
391 |
by Auto_tac; |
|
392 |
qed "hypreal_of_real_not_eq_omega"; |
|
393 |
||
394 |
(* existence of infinitesimal number also not *) |
|
395 |
(* corresponding to any real number *) |
|
396 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
397 |
Goal "inverse (real (x::nat)) = inverse (real y) ==> x = y"; |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
398 |
by (rtac (inj_real_of_nat RS injD) 1); |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
399 |
by (Asm_full_simp_tac 1); |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
400 |
qed "real_of_nat_inverse_inj"; |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
401 |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
402 |
Goal "{n::nat. x = inverse(real(Suc n))} = {} | \ |
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
403 |
\ (EX y. {n::nat. x = inverse(real(Suc n))} = {y})"; |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
404 |
by (auto_tac (claset(), simpset() addsimps [inj_real_of_nat RS inj_eq])); |
10751 | 405 |
qed "lemma_epsilon_empty_singleton_disj"; |
406 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
407 |
Goal "finite {n. x = inverse(real(Suc n))}"; |
10751 | 408 |
by (cut_inst_tac [("x","x")] lemma_epsilon_empty_singleton_disj 1); |
409 |
by Auto_tac; |
|
410 |
qed "lemma_finite_epsilon_set"; |
|
411 |
||
412 |
Goalw [epsilon_def,hypreal_of_real_def] |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
413 |
"~ (EX x. hypreal_of_real x = epsilon)"; |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
414 |
by (auto_tac (claset(), |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
415 |
simpset() addsimps [lemma_finite_epsilon_set RS FreeUltrafilterNat_finite])); |
10751 | 416 |
qed "not_ex_hypreal_of_real_eq_epsilon"; |
417 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
418 |
Goal "hypreal_of_real x ~= epsilon"; |
10751 | 419 |
by (cut_facts_tac [not_ex_hypreal_of_real_eq_epsilon] 1); |
420 |
by Auto_tac; |
|
421 |
qed "hypreal_of_real_not_eq_epsilon"; |
|
422 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
423 |
Goalw [epsilon_def,hypreal_zero_def] "epsilon ~= 0"; |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
424 |
by Auto_tac; |
10751 | 425 |
by (auto_tac (claset(), |
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
426 |
simpset() addsimps [real_of_nat_Suc_gt_zero RS real_not_refl2 RS not_sym])); |
10751 | 427 |
qed "hypreal_epsilon_not_zero"; |
428 |
||
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10784
diff
changeset
|
429 |
Goal "epsilon = inverse(omega)"; |
10751 | 430 |
by (asm_full_simp_tac (simpset() addsimps |
431 |
[hypreal_inverse, omega_def, epsilon_def]) 1); |
|
432 |
qed "hypreal_epsilon_inverse_omega"; |
|
433 |
||
434 |
||
435 |
(* this proof is so much simpler than one for reals!! *) |
|
436 |
Goal "[| 0 < r; r < x |] ==> inverse x < inverse (r::hypreal)"; |
|
437 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
438 |
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1); |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
439 |
by (auto_tac (claset(), |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
440 |
simpset() addsimps [hypreal_inverse, hypreal_less,hypreal_zero_def])); |
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
441 |
by (ultra_tac (claset() addIs [real_inverse_less_swap], simpset()) 1); |
10751 | 442 |
qed "hypreal_inverse_less_swap"; |
443 |
||
444 |
Goal "[| 0 < r; 0 < x|] ==> (r < x) = (inverse x < inverse (r::hypreal))"; |
|
10778
2c6605049646
more tidying, especially to remove real_of_posnat
paulson
parents:
10751
diff
changeset
|
445 |
by (auto_tac (claset() addIs [hypreal_inverse_less_swap], simpset())); |
10751 | 446 |
by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1); |
447 |
by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1); |
|
448 |
by (rtac hypreal_inverse_less_swap 1); |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
449 |
by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_gt_0])); |
10751 | 450 |
qed "hypreal_inverse_less_iff"; |
451 |
||
452 |
Goal "[| 0 < z; x < y |] ==> x * inverse z < y * inverse (z::hypreal)"; |
|
453 |
by (blast_tac (claset() addSIs [hypreal_mult_less_mono1, |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
454 |
hypreal_inverse_gt_0]) 1); |
10751 | 455 |
qed "hypreal_mult_inverse_less_mono1"; |
456 |
||
457 |
Goal "[| 0 < z; x < y |] ==> inverse z * x < inverse (z::hypreal) * y"; |
|
458 |
by (blast_tac (claset() addSIs [hypreal_mult_less_mono2, |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
459 |
hypreal_inverse_gt_0]) 1); |
10751 | 460 |
qed "hypreal_mult_inverse_less_mono2"; |
461 |
||
462 |
Goal "[| (0::hypreal) < z; x*z < y*z |] ==> x < y"; |
|
463 |
by (forw_inst_tac [("x","x*z")] hypreal_mult_inverse_less_mono1 1); |
|
464 |
by (dtac (hypreal_not_refl2 RS not_sym) 2); |
|
465 |
by (auto_tac (claset() addSDs [hypreal_mult_inverse], |
|
466 |
simpset() addsimps hypreal_mult_ac)); |
|
467 |
qed "hypreal_less_mult_right_cancel"; |
|
468 |
||
469 |
Goal "[| (0::hypreal) < z; z*x < z*y |] ==> x < y"; |
|
470 |
by (auto_tac (claset() addIs [hypreal_less_mult_right_cancel], |
|
471 |
simpset() addsimps [hypreal_mult_commute])); |
|
472 |
qed "hypreal_less_mult_left_cancel"; |
|
473 |
||
474 |
Goal "[| 0 < r; (0::hypreal) < ra; r < x; ra < y |] ==> r*ra < x*y"; |
|
475 |
by (forw_inst_tac [("y","r")] order_less_trans 1); |
|
476 |
by (dres_inst_tac [("z","ra"),("x","r")] hypreal_mult_less_mono1 2); |
|
477 |
by (dres_inst_tac [("z","x"),("x","ra")] hypreal_mult_less_mono2 3); |
|
478 |
by (auto_tac (claset() addIs [order_less_trans], simpset())); |
|
479 |
qed "hypreal_mult_less_gt_zero"; |
|
480 |
||
481 |
Goal "[| 0 < r; (0::hypreal) < ra; r <= x; ra <= y |] ==> r*ra <= x*y"; |
|
482 |
by (REPEAT(dtac order_le_imp_less_or_eq 1)); |
|
483 |
by (rtac hypreal_less_or_eq_imp_le 1); |
|
484 |
by (auto_tac (claset() addIs [hypreal_mult_less_mono1, |
|
485 |
hypreal_mult_less_mono2,hypreal_mult_less_gt_zero], |
|
486 |
simpset())); |
|
487 |
qed "hypreal_mult_le_ge_zero"; |
|
488 |
||
489 |
(*---------------------------------------------------------------------------- |
|
490 |
Some convenient biconditionals for products of signs |
|
491 |
----------------------------------------------------------------------------*) |
|
492 |
||
493 |
Goal "((0::hypreal) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)"; |
|
494 |
by (auto_tac (claset(), |
|
495 |
simpset() addsimps [order_le_less, linorder_not_less, |
|
496 |
hypreal_mult_order, hypreal_mult_less_zero1])); |
|
497 |
by (ALLGOALS (rtac ccontr)); |
|
498 |
by (auto_tac (claset(), |
|
499 |
simpset() addsimps [order_le_less, linorder_not_less])); |
|
500 |
by (ALLGOALS (etac rev_mp)); |
|
501 |
by (ALLGOALS (dtac hypreal_mult_less_zero THEN' assume_tac)); |
|
502 |
by (auto_tac (claset() addDs [order_less_not_sym], |
|
503 |
simpset() addsimps [hypreal_mult_commute])); |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
504 |
qed "hypreal_0_less_mult_iff"; |
10751 | 505 |
|
506 |
Goal "((0::hypreal) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)"; |
|
507 |
by (auto_tac (claset() addDs [hypreal_mult_zero_disj], |
|
508 |
simpset() addsimps [order_le_less, linorder_not_less, |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
509 |
hypreal_0_less_mult_iff])); |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
510 |
qed "hypreal_0_le_mult_iff"; |
10751 | 511 |
|
512 |
Goal "(x*y < (0::hypreal)) = (0 < x & y < 0 | x < 0 & 0 < y)"; |
|
513 |
by (auto_tac (claset(), |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
514 |
simpset() addsimps [hypreal_0_le_mult_iff, |
10751 | 515 |
linorder_not_le RS sym])); |
516 |
by (auto_tac (claset() addDs [order_less_not_sym], |
|
517 |
simpset() addsimps [linorder_not_le])); |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
518 |
qed "hypreal_mult_less_0_iff"; |
10751 | 519 |
|
520 |
Goal "(x*y <= (0::hypreal)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)"; |
|
521 |
by (auto_tac (claset() addDs [order_less_not_sym], |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
522 |
simpset() addsimps [hypreal_0_less_mult_iff, |
10751 | 523 |
linorder_not_less RS sym])); |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
524 |
qed "hypreal_mult_le_0_iff"; |
10751 | 525 |