author | paulson |
Tue, 19 Dec 2000 15:16:21 +0100 | |
changeset 10702 | 9e6898befad4 |
parent 10690 | cd80241125b0 |
child 10712 | 351ba950d4d9 |
permissions | -rw-r--r-- |
7218 | 1 |
(* Title : HOL/Real/Hyperreal/Hyper.ML |
2 |
ID : $Id$ |
|
3 |
Author : Jacques D. Fleuriot |
|
4 |
Copyright : 1998 University of Cambridge |
|
5 |
Description : Ultrapower construction of hyperreals |
|
6 |
*) |
|
7 |
||
8 |
(*------------------------------------------------------------------------ |
|
9 |
Proof that the set of naturals is not finite |
|
10 |
------------------------------------------------------------------------*) |
|
11 |
||
12 |
(*** based on James' proof that the set of naturals is not finite ***) |
|
9055 | 13 |
Goal "finite (A::nat set) --> (EX n. ALL m. Suc (n + m) ~: A)"; |
7218 | 14 |
by (rtac impI 1); |
15 |
by (eres_inst_tac [("F","A")] finite_induct 1); |
|
16 |
by (Blast_tac 1 THEN etac exE 1); |
|
17 |
by (res_inst_tac [("x","n + x")] exI 1); |
|
18 |
by (rtac allI 1 THEN eres_inst_tac [("x","x + m")] allE 1); |
|
19 |
by (auto_tac (claset(), simpset() addsimps add_ac)); |
|
20 |
by (auto_tac (claset(), |
|
21 |
simpset() addsimps [add_assoc RS sym, |
|
22 |
less_add_Suc2 RS less_not_refl2])); |
|
23 |
qed_spec_mp "finite_exhausts"; |
|
24 |
||
9055 | 25 |
Goal "finite (A :: nat set) --> (EX n. n ~:A)"; |
7218 | 26 |
by (rtac impI 1 THEN dtac finite_exhausts 1); |
27 |
by (Blast_tac 1); |
|
28 |
qed_spec_mp "finite_not_covers"; |
|
29 |
||
30 |
Goal "~ finite(UNIV:: nat set)"; |
|
31 |
by (fast_tac (claset() addSDs [finite_exhausts]) 1); |
|
32 |
qed "not_finite_nat"; |
|
33 |
||
34 |
(*------------------------------------------------------------------------ |
|
35 |
Existence of free ultrafilter over the naturals and proof of various |
|
36 |
properties of the FreeUltrafilterNat- an arbitrary free ultrafilter |
|
37 |
------------------------------------------------------------------------*) |
|
38 |
||
39 |
Goal "EX U. U: FreeUltrafilter (UNIV::nat set)"; |
|
40 |
by (rtac (not_finite_nat RS FreeUltrafilter_Ex) 1); |
|
41 |
qed "FreeUltrafilterNat_Ex"; |
|
42 |
||
43 |
Goalw [FreeUltrafilterNat_def] |
|
44 |
"FreeUltrafilterNat: FreeUltrafilter(UNIV:: nat set)"; |
|
45 |
by (rtac (FreeUltrafilterNat_Ex RS exE) 1); |
|
9969 | 46 |
by (rtac someI2 1 THEN ALLGOALS(assume_tac)); |
7218 | 47 |
qed "FreeUltrafilterNat_mem"; |
48 |
Addsimps [FreeUltrafilterNat_mem]; |
|
49 |
||
50 |
Goalw [FreeUltrafilterNat_def] "finite x ==> x ~: FreeUltrafilterNat"; |
|
51 |
by (rtac (FreeUltrafilterNat_Ex RS exE) 1); |
|
9969 | 52 |
by (rtac someI2 1 THEN assume_tac 1); |
7218 | 53 |
by (blast_tac (claset() addDs [mem_FreeUltrafiltersetD1]) 1); |
54 |
qed "FreeUltrafilterNat_finite"; |
|
55 |
||
56 |
Goal "x: FreeUltrafilterNat ==> ~ finite x"; |
|
57 |
by (blast_tac (claset() addDs [FreeUltrafilterNat_finite]) 1); |
|
58 |
qed "FreeUltrafilterNat_not_finite"; |
|
59 |
||
60 |
Goalw [FreeUltrafilterNat_def] "{} ~: FreeUltrafilterNat"; |
|
61 |
by (rtac (FreeUltrafilterNat_Ex RS exE) 1); |
|
9969 | 62 |
by (rtac someI2 1 THEN assume_tac 1); |
7218 | 63 |
by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter, |
64 |
Ultrafilter_Filter,Filter_empty_not_mem]) 1); |
|
65 |
qed "FreeUltrafilterNat_empty"; |
|
66 |
Addsimps [FreeUltrafilterNat_empty]; |
|
67 |
||
68 |
Goal "[| X: FreeUltrafilterNat; Y: FreeUltrafilterNat |] \ |
|
69 |
\ ==> X Int Y : FreeUltrafilterNat"; |
|
70 |
by (cut_facts_tac [FreeUltrafilterNat_mem] 1); |
|
71 |
by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter, |
|
72 |
Ultrafilter_Filter,mem_FiltersetD1]) 1); |
|
73 |
qed "FreeUltrafilterNat_Int"; |
|
74 |
||
75 |
Goal "[| X: FreeUltrafilterNat; X <= Y |] \ |
|
76 |
\ ==> Y : FreeUltrafilterNat"; |
|
77 |
by (cut_facts_tac [FreeUltrafilterNat_mem] 1); |
|
78 |
by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter, |
|
79 |
Ultrafilter_Filter,mem_FiltersetD2]) 1); |
|
80 |
qed "FreeUltrafilterNat_subset"; |
|
81 |
||
82 |
Goal "X: FreeUltrafilterNat ==> -X ~: FreeUltrafilterNat"; |
|
83 |
by (Step_tac 1); |
|
84 |
by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1); |
|
85 |
by Auto_tac; |
|
86 |
qed "FreeUltrafilterNat_Compl"; |
|
87 |
||
88 |
Goal "X~: FreeUltrafilterNat ==> -X : FreeUltrafilterNat"; |
|
89 |
by (cut_facts_tac [FreeUltrafilterNat_mem RS (FreeUltrafilter_iff RS iffD1)] 1); |
|
90 |
by (Step_tac 1 THEN dres_inst_tac [("x","X")] bspec 1); |
|
10677 | 91 |
by (auto_tac (claset(), simpset() addsimps [UNIV_diff_Compl])); |
7218 | 92 |
qed "FreeUltrafilterNat_Compl_mem"; |
93 |
||
94 |
Goal "(X ~: FreeUltrafilterNat) = (-X: FreeUltrafilterNat)"; |
|
95 |
by (blast_tac (claset() addDs [FreeUltrafilterNat_Compl, |
|
96 |
FreeUltrafilterNat_Compl_mem]) 1); |
|
97 |
qed "FreeUltrafilterNat_Compl_iff1"; |
|
98 |
||
99 |
Goal "(X: FreeUltrafilterNat) = (-X ~: FreeUltrafilterNat)"; |
|
100 |
by (auto_tac (claset(), |
|
101 |
simpset() addsimps [FreeUltrafilterNat_Compl_iff1 RS sym])); |
|
102 |
qed "FreeUltrafilterNat_Compl_iff2"; |
|
103 |
||
104 |
Goal "(UNIV::nat set) : FreeUltrafilterNat"; |
|
105 |
by (rtac (FreeUltrafilterNat_mem RS FreeUltrafilter_Ultrafilter RS |
|
106 |
Ultrafilter_Filter RS mem_FiltersetD4) 1); |
|
107 |
qed "FreeUltrafilterNat_UNIV"; |
|
108 |
Addsimps [FreeUltrafilterNat_UNIV]; |
|
109 |
||
9391 | 110 |
Goal "UNIV : FreeUltrafilterNat"; |
7218 | 111 |
by Auto_tac; |
112 |
qed "FreeUltrafilterNat_Nat_set"; |
|
113 |
Addsimps [FreeUltrafilterNat_Nat_set]; |
|
114 |
||
115 |
Goal "{n. P(n) = P(n)} : FreeUltrafilterNat"; |
|
116 |
by (Simp_tac 1); |
|
117 |
qed "FreeUltrafilterNat_Nat_set_refl"; |
|
118 |
AddIs [FreeUltrafilterNat_Nat_set_refl]; |
|
119 |
||
120 |
Goal "{n::nat. P} : FreeUltrafilterNat ==> P"; |
|
121 |
by (rtac ccontr 1); |
|
122 |
by (rotate_tac 1 1); |
|
123 |
by (Asm_full_simp_tac 1); |
|
124 |
qed "FreeUltrafilterNat_P"; |
|
125 |
||
126 |
Goal "{n. P(n)} : FreeUltrafilterNat ==> EX n. P(n)"; |
|
127 |
by (rtac ccontr 1 THEN rotate_tac 1 1); |
|
128 |
by (Asm_full_simp_tac 1); |
|
129 |
qed "FreeUltrafilterNat_Ex_P"; |
|
130 |
||
131 |
Goal "ALL n. P(n) ==> {n. P(n)} : FreeUltrafilterNat"; |
|
10677 | 132 |
by (auto_tac (claset() addIs [FreeUltrafilterNat_Nat_set], simpset())); |
7218 | 133 |
qed "FreeUltrafilterNat_all"; |
134 |
||
10043 | 135 |
(*------------------------------------------------------- |
7218 | 136 |
Define and use Ultrafilter tactics |
10043 | 137 |
-------------------------------------------------------*) |
7218 | 138 |
use "fuf.ML"; |
139 |
||
10043 | 140 |
(*------------------------------------------------------- |
141 |
Now prove one further property of our free ultrafilter |
|
7218 | 142 |
-------------------------------------------------------*) |
143 |
Goal "X Un Y: FreeUltrafilterNat \ |
|
144 |
\ ==> X: FreeUltrafilterNat | Y: FreeUltrafilterNat"; |
|
145 |
by Auto_tac; |
|
146 |
by (Ultra_tac 1); |
|
147 |
qed "FreeUltrafilterNat_Un"; |
|
148 |
||
10043 | 149 |
(*------------------------------------------------------- |
150 |
Properties of hyprel |
|
151 |
-------------------------------------------------------*) |
|
7218 | 152 |
|
153 |
(** Proving that hyprel is an equivalence relation **) |
|
154 |
(** Natural deduction for hyprel **) |
|
155 |
||
156 |
Goalw [hyprel_def] |
|
157 |
"((X,Y): hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)"; |
|
158 |
by (Fast_tac 1); |
|
159 |
qed "hyprel_iff"; |
|
160 |
||
161 |
Goalw [hyprel_def] |
|
162 |
"{n. X n = Y n}: FreeUltrafilterNat ==> (X,Y): hyprel"; |
|
163 |
by (Fast_tac 1); |
|
164 |
qed "hyprelI"; |
|
165 |
||
166 |
Goalw [hyprel_def] |
|
167 |
"p: hyprel --> (EX X Y. \ |
|
168 |
\ p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)"; |
|
169 |
by (Fast_tac 1); |
|
170 |
qed "hyprelE_lemma"; |
|
171 |
||
9432 | 172 |
val [major,minor] = goal (the_context ()) |
7218 | 173 |
"[| p: hyprel; \ |
174 |
\ !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\ |
|
175 |
\ |] ==> Q |] ==> Q"; |
|
176 |
by (cut_facts_tac [major RS (hyprelE_lemma RS mp)] 1); |
|
177 |
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); |
|
178 |
qed "hyprelE"; |
|
179 |
||
180 |
AddSIs [hyprelI]; |
|
181 |
AddSEs [hyprelE]; |
|
182 |
||
183 |
Goalw [hyprel_def] "(x,x): hyprel"; |
|
10677 | 184 |
by (auto_tac (claset(), |
185 |
simpset() addsimps [FreeUltrafilterNat_Nat_set])); |
|
7218 | 186 |
qed "hyprel_refl"; |
187 |
||
188 |
Goal "{n. X n = Y n} = {n. Y n = X n}"; |
|
189 |
by Auto_tac; |
|
190 |
qed "lemma_perm"; |
|
191 |
||
192 |
Goalw [hyprel_def] "(x,y): hyprel --> (y,x):hyprel"; |
|
10677 | 193 |
by (auto_tac (claset() addIs [lemma_perm RS subst], simpset())); |
7218 | 194 |
qed_spec_mp "hyprel_sym"; |
195 |
||
196 |
Goalw [hyprel_def] |
|
197 |
"(x,y): hyprel --> (y,z):hyprel --> (x,z):hyprel"; |
|
198 |
by Auto_tac; |
|
199 |
by (Ultra_tac 1); |
|
200 |
qed_spec_mp "hyprel_trans"; |
|
201 |
||
9391 | 202 |
Goalw [equiv_def, refl_def, sym_def, trans_def] "equiv UNIV hyprel"; |
7218 | 203 |
by (auto_tac (claset() addSIs [hyprel_refl] |
204 |
addSEs [hyprel_sym,hyprel_trans] |
|
205 |
delrules [hyprelI,hyprelE], |
|
206 |
simpset() addsimps [FreeUltrafilterNat_Nat_set])); |
|
207 |
qed "equiv_hyprel"; |
|
208 |
||
9391 | 209 |
(* (hyprel ^^ {x} = hyprel ^^ {y}) = ((x,y) : hyprel) *) |
9108 | 210 |
bind_thm ("equiv_hyprel_iff", |
9391 | 211 |
[equiv_hyprel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff); |
7218 | 212 |
|
213 |
Goalw [hypreal_def,hyprel_def,quotient_def] "hyprel^^{x}:hypreal"; |
|
214 |
by (Blast_tac 1); |
|
215 |
qed "hyprel_in_hypreal"; |
|
216 |
||
217 |
Goal "inj_on Abs_hypreal hypreal"; |
|
218 |
by (rtac inj_on_inverseI 1); |
|
219 |
by (etac Abs_hypreal_inverse 1); |
|
220 |
qed "inj_on_Abs_hypreal"; |
|
221 |
||
222 |
Addsimps [equiv_hyprel_iff,inj_on_Abs_hypreal RS inj_on_iff, |
|
223 |
hyprel_iff, hyprel_in_hypreal, Abs_hypreal_inverse]; |
|
224 |
||
225 |
Addsimps [equiv_hyprel RS eq_equiv_class_iff]; |
|
9108 | 226 |
bind_thm ("eq_hyprelD", equiv_hyprel RSN (2,eq_equiv_class)); |
7218 | 227 |
|
228 |
Goal "inj(Rep_hypreal)"; |
|
229 |
by (rtac inj_inverseI 1); |
|
230 |
by (rtac Rep_hypreal_inverse 1); |
|
231 |
qed "inj_Rep_hypreal"; |
|
232 |
||
233 |
Goalw [hyprel_def] "x: hyprel ^^ {x}"; |
|
234 |
by (Step_tac 1); |
|
10677 | 235 |
by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset())); |
7218 | 236 |
qed "lemma_hyprel_refl"; |
237 |
||
238 |
Addsimps [lemma_hyprel_refl]; |
|
239 |
||
240 |
Goalw [hypreal_def] "{} ~: hypreal"; |
|
241 |
by (auto_tac (claset() addSEs [quotientE], simpset())); |
|
242 |
qed "hypreal_empty_not_mem"; |
|
243 |
||
244 |
Addsimps [hypreal_empty_not_mem]; |
|
245 |
||
246 |
Goal "Rep_hypreal x ~= {}"; |
|
247 |
by (cut_inst_tac [("x","x")] Rep_hypreal 1); |
|
248 |
by Auto_tac; |
|
249 |
qed "Rep_hypreal_nonempty"; |
|
250 |
||
251 |
Addsimps [Rep_hypreal_nonempty]; |
|
252 |
||
253 |
(*------------------------------------------------------------------------ |
|
254 |
hypreal_of_real: the injection from real to hypreal |
|
255 |
------------------------------------------------------------------------*) |
|
256 |
||
257 |
Goal "inj(hypreal_of_real)"; |
|
258 |
by (rtac injI 1); |
|
259 |
by (rewtac hypreal_of_real_def); |
|
260 |
by (dtac (inj_on_Abs_hypreal RS inj_onD) 1); |
|
261 |
by (REPEAT (rtac hyprel_in_hypreal 1)); |
|
262 |
by (dtac eq_equiv_class 1); |
|
263 |
by (rtac equiv_hyprel 1); |
|
264 |
by (Fast_tac 1); |
|
265 |
by (rtac ccontr 1 THEN rotate_tac 1 1); |
|
266 |
by Auto_tac; |
|
267 |
qed "inj_hypreal_of_real"; |
|
268 |
||
9432 | 269 |
val [prem] = goal (the_context ()) |
7218 | 270 |
"(!!x y. z = Abs_hypreal(hyprel^^{x}) ==> P) ==> P"; |
271 |
by (res_inst_tac [("x1","z")] |
|
272 |
(rewrite_rule [hypreal_def] Rep_hypreal RS quotientE) 1); |
|
273 |
by (dres_inst_tac [("f","Abs_hypreal")] arg_cong 1); |
|
274 |
by (res_inst_tac [("x","x")] prem 1); |
|
275 |
by (asm_full_simp_tac (simpset() addsimps [Rep_hypreal_inverse]) 1); |
|
276 |
qed "eq_Abs_hypreal"; |
|
277 |
||
278 |
(**** hypreal_minus: additive inverse on hypreal ****) |
|
279 |
||
280 |
Goalw [congruent_def] |
|
281 |
"congruent hyprel (%X. hyprel^^{%n. - (X n)})"; |
|
282 |
by Safe_tac; |
|
283 |
by (ALLGOALS Ultra_tac); |
|
284 |
qed "hypreal_minus_congruent"; |
|
285 |
||
286 |
Goalw [hypreal_minus_def] |
|
10677 | 287 |
"- (Abs_hypreal(hyprel^^{%n. X n})) = Abs_hypreal(hyprel ^^ {%n. -(X n)})"; |
7218 | 288 |
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); |
289 |
by (simp_tac (simpset() addsimps |
|
10677 | 290 |
[hyprel_in_hypreal RS Abs_hypreal_inverse, |
291 |
[equiv_hyprel, hypreal_minus_congruent] MRS UN_equiv_class]) 1); |
|
7218 | 292 |
qed "hypreal_minus"; |
293 |
||
294 |
Goal "- (- z) = (z::hypreal)"; |
|
295 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
|
296 |
by (asm_simp_tac (simpset() addsimps [hypreal_minus]) 1); |
|
297 |
qed "hypreal_minus_minus"; |
|
298 |
||
299 |
Addsimps [hypreal_minus_minus]; |
|
300 |
||
301 |
Goal "inj(%r::hypreal. -r)"; |
|
302 |
by (rtac injI 1); |
|
303 |
by (dres_inst_tac [("f","uminus")] arg_cong 1); |
|
304 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_minus_minus]) 1); |
|
305 |
qed "inj_hypreal_minus"; |
|
306 |
||
9055 | 307 |
Goalw [hypreal_zero_def] "-0 = (0::hypreal)"; |
7218 | 308 |
by (simp_tac (simpset() addsimps [hypreal_minus]) 1); |
309 |
qed "hypreal_minus_zero"; |
|
310 |
Addsimps [hypreal_minus_zero]; |
|
311 |
||
9055 | 312 |
Goal "(-x = 0) = (x = (0::hypreal))"; |
7218 | 313 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
10677 | 314 |
by (auto_tac (claset(), |
315 |
simpset() addsimps [hypreal_zero_def, hypreal_minus, eq_commute] @ |
|
316 |
real_add_ac)); |
|
7218 | 317 |
qed "hypreal_minus_zero_iff"; |
318 |
||
319 |
Addsimps [hypreal_minus_zero_iff]; |
|
320 |
||
321 |
||
322 |
(**** hyperreal addition: hypreal_add ****) |
|
323 |
||
324 |
Goalw [congruent2_def] |
|
325 |
"congruent2 hyprel (%X Y. hyprel^^{%n. X n + Y n})"; |
|
326 |
by Safe_tac; |
|
327 |
by (ALLGOALS(Ultra_tac)); |
|
328 |
qed "hypreal_add_congruent2"; |
|
329 |
||
330 |
Goalw [hypreal_add_def] |
|
331 |
"Abs_hypreal(hyprel^^{%n. X n}) + Abs_hypreal(hyprel^^{%n. Y n}) = \ |
|
332 |
\ Abs_hypreal(hyprel^^{%n. X n + Y n})"; |
|
9391 | 333 |
by (simp_tac (simpset() addsimps |
334 |
[[equiv_hyprel, hypreal_add_congruent2] MRS UN_equiv_class2]) 1); |
|
7218 | 335 |
qed "hypreal_add"; |
336 |
||
337 |
Goal "(z::hypreal) + w = w + z"; |
|
338 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
|
339 |
by (res_inst_tac [("z","w")] eq_Abs_hypreal 1); |
|
340 |
by (asm_simp_tac (simpset() addsimps (real_add_ac @ [hypreal_add])) 1); |
|
341 |
qed "hypreal_add_commute"; |
|
342 |
||
343 |
Goal "((z1::hypreal) + z2) + z3 = z1 + (z2 + z3)"; |
|
344 |
by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1); |
|
345 |
by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1); |
|
346 |
by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1); |
|
347 |
by (asm_simp_tac (simpset() addsimps [hypreal_add, real_add_assoc]) 1); |
|
348 |
qed "hypreal_add_assoc"; |
|
349 |
||
350 |
(*For AC rewriting*) |
|
351 |
Goal "(x::hypreal)+(y+z)=y+(x+z)"; |
|
352 |
by (rtac (hypreal_add_commute RS trans) 1); |
|
353 |
by (rtac (hypreal_add_assoc RS trans) 1); |
|
354 |
by (rtac (hypreal_add_commute RS arg_cong) 1); |
|
355 |
qed "hypreal_add_left_commute"; |
|
356 |
||
357 |
(* hypreal addition is an AC operator *) |
|
9108 | 358 |
bind_thms ("hypreal_add_ac", [hypreal_add_assoc,hypreal_add_commute, |
359 |
hypreal_add_left_commute]); |
|
7218 | 360 |
|
9055 | 361 |
Goalw [hypreal_zero_def] "(0::hypreal) + z = z"; |
7218 | 362 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
363 |
by (asm_full_simp_tac (simpset() addsimps |
|
364 |
[hypreal_add]) 1); |
|
365 |
qed "hypreal_add_zero_left"; |
|
366 |
||
9055 | 367 |
Goal "z + (0::hypreal) = z"; |
7218 | 368 |
by (simp_tac (simpset() addsimps |
369 |
[hypreal_add_zero_left,hypreal_add_commute]) 1); |
|
370 |
qed "hypreal_add_zero_right"; |
|
371 |
||
9055 | 372 |
Goalw [hypreal_zero_def] "z + -z = (0::hypreal)"; |
7218 | 373 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
10648 | 374 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_minus, hypreal_add]) 1); |
7218 | 375 |
qed "hypreal_add_minus"; |
376 |
||
9055 | 377 |
Goal "-z + z = (0::hypreal)"; |
10648 | 378 |
by (simp_tac (simpset() addsimps [hypreal_add_commute, hypreal_add_minus]) 1); |
7218 | 379 |
qed "hypreal_add_minus_left"; |
380 |
||
381 |
Addsimps [hypreal_add_minus,hypreal_add_minus_left, |
|
382 |
hypreal_add_zero_left,hypreal_add_zero_right]; |
|
383 |
||
9055 | 384 |
Goal "EX y. (x::hypreal) + y = 0"; |
7218 | 385 |
by (fast_tac (claset() addIs [hypreal_add_minus]) 1); |
386 |
qed "hypreal_minus_ex"; |
|
387 |
||
9055 | 388 |
Goal "EX! y. (x::hypreal) + y = 0"; |
10677 | 389 |
by (auto_tac (claset() addIs [hypreal_add_minus], simpset())); |
7218 | 390 |
by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1); |
391 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
|
392 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
|
393 |
qed "hypreal_minus_ex1"; |
|
394 |
||
9055 | 395 |
Goal "EX! y. y + (x::hypreal) = 0"; |
10677 | 396 |
by (auto_tac (claset() addIs [hypreal_add_minus_left], simpset())); |
7218 | 397 |
by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1); |
398 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1); |
|
399 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
|
400 |
qed "hypreal_minus_left_ex1"; |
|
401 |
||
9055 | 402 |
Goal "x + y = (0::hypreal) ==> x = -y"; |
7218 | 403 |
by (cut_inst_tac [("z","y")] hypreal_add_minus_left 1); |
404 |
by (res_inst_tac [("x1","y")] (hypreal_minus_left_ex1 RS ex1E) 1); |
|
405 |
by (Blast_tac 1); |
|
406 |
qed "hypreal_add_minus_eq_minus"; |
|
407 |
||
9055 | 408 |
Goal "EX y::hypreal. x = -y"; |
7218 | 409 |
by (cut_inst_tac [("x","x")] hypreal_minus_ex 1); |
410 |
by (etac exE 1 THEN dtac hypreal_add_minus_eq_minus 1); |
|
411 |
by (Fast_tac 1); |
|
412 |
qed "hypreal_as_add_inverse_ex"; |
|
413 |
||
414 |
Goal "-(x + (y::hypreal)) = -x + -y"; |
|
415 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
416 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
10648 | 417 |
by (auto_tac (claset(), |
418 |
simpset() addsimps [hypreal_minus, hypreal_add, |
|
419 |
real_minus_add_distrib])); |
|
7218 | 420 |
qed "hypreal_minus_add_distrib"; |
10043 | 421 |
Addsimps [hypreal_minus_add_distrib]; |
7218 | 422 |
|
423 |
Goal "-(y + -(x::hypreal)) = x + -y"; |
|
10043 | 424 |
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
7218 | 425 |
qed "hypreal_minus_distrib1"; |
426 |
||
427 |
Goal "(x + - (y::hypreal)) + (y + - z) = x + -z"; |
|
428 |
by (res_inst_tac [("w1","y")] (hypreal_add_commute RS subst) 1); |
|
429 |
by (simp_tac (simpset() addsimps [hypreal_add_left_commute, |
|
10648 | 430 |
hypreal_add_assoc]) 1); |
7218 | 431 |
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
432 |
qed "hypreal_add_minus_cancel1"; |
|
433 |
||
434 |
Goal "((x::hypreal) + y = x + z) = (y = z)"; |
|
435 |
by (Step_tac 1); |
|
436 |
by (dres_inst_tac [("f","%t.-x + t")] arg_cong 1); |
|
437 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
|
438 |
qed "hypreal_add_left_cancel"; |
|
439 |
||
440 |
Goal "z + (x + (y + -z)) = x + (y::hypreal)"; |
|
441 |
by (simp_tac (simpset() addsimps hypreal_add_ac) 1); |
|
442 |
qed "hypreal_add_minus_cancel2"; |
|
443 |
Addsimps [hypreal_add_minus_cancel2]; |
|
444 |
||
445 |
Goal "y + -(x + y) = -(x::hypreal)"; |
|
10043 | 446 |
by (Full_simp_tac 1); |
7218 | 447 |
by (rtac (hypreal_add_left_commute RS subst) 1); |
448 |
by (Full_simp_tac 1); |
|
449 |
qed "hypreal_add_minus_cancel"; |
|
450 |
Addsimps [hypreal_add_minus_cancel]; |
|
451 |
||
452 |
Goal "y + -(y + x) = -(x::hypreal)"; |
|
10043 | 453 |
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
7218 | 454 |
qed "hypreal_add_minus_cancelc"; |
455 |
Addsimps [hypreal_add_minus_cancelc]; |
|
456 |
||
457 |
Goal "(z + -x) + (y + -z) = (y + -(x::hypreal))"; |
|
10648 | 458 |
by (full_simp_tac |
459 |
(simpset() addsimps [hypreal_minus_add_distrib RS sym, |
|
460 |
hypreal_add_left_cancel] @ hypreal_add_ac |
|
461 |
delsimps [hypreal_minus_add_distrib]) 1); |
|
7218 | 462 |
qed "hypreal_add_minus_cancel3"; |
463 |
Addsimps [hypreal_add_minus_cancel3]; |
|
464 |
||
465 |
Goal "(y + (x::hypreal)= z + x) = (y = z)"; |
|
466 |
by (simp_tac (simpset() addsimps [hypreal_add_commute, |
|
10648 | 467 |
hypreal_add_left_cancel]) 1); |
7218 | 468 |
qed "hypreal_add_right_cancel"; |
469 |
||
470 |
Goal "z + (y + -z) = (y::hypreal)"; |
|
471 |
by (simp_tac (simpset() addsimps hypreal_add_ac) 1); |
|
472 |
qed "hypreal_add_minus_cancel4"; |
|
473 |
Addsimps [hypreal_add_minus_cancel4]; |
|
474 |
||
475 |
Goal "z + (w + (x + (-z + y))) = w + x + (y::hypreal)"; |
|
476 |
by (simp_tac (simpset() addsimps hypreal_add_ac) 1); |
|
477 |
qed "hypreal_add_minus_cancel5"; |
|
478 |
Addsimps [hypreal_add_minus_cancel5]; |
|
479 |
||
10043 | 480 |
Goal "z + ((- z) + w) = (w::hypreal)"; |
481 |
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
|
482 |
qed "hypreal_add_minus_cancelA"; |
|
483 |
||
484 |
Goal "(-z) + (z + w) = (w::hypreal)"; |
|
485 |
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
|
486 |
qed "hypreal_minus_add_cancelA"; |
|
487 |
||
488 |
Addsimps [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA]; |
|
7218 | 489 |
|
490 |
(**** hyperreal multiplication: hypreal_mult ****) |
|
491 |
||
492 |
Goalw [congruent2_def] |
|
493 |
"congruent2 hyprel (%X Y. hyprel^^{%n. X n * Y n})"; |
|
494 |
by Safe_tac; |
|
495 |
by (ALLGOALS(Ultra_tac)); |
|
496 |
qed "hypreal_mult_congruent2"; |
|
497 |
||
498 |
Goalw [hypreal_mult_def] |
|
499 |
"Abs_hypreal(hyprel^^{%n. X n}) * Abs_hypreal(hyprel^^{%n. Y n}) = \ |
|
500 |
\ Abs_hypreal(hyprel^^{%n. X n * Y n})"; |
|
9391 | 501 |
by (simp_tac (simpset() addsimps |
502 |
[[equiv_hyprel, hypreal_mult_congruent2] MRS UN_equiv_class2]) 1); |
|
7218 | 503 |
qed "hypreal_mult"; |
504 |
||
505 |
Goal "(z::hypreal) * w = w * z"; |
|
506 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
|
507 |
by (res_inst_tac [("z","w")] eq_Abs_hypreal 1); |
|
508 |
by (asm_simp_tac (simpset() addsimps ([hypreal_mult] @ real_mult_ac)) 1); |
|
509 |
qed "hypreal_mult_commute"; |
|
510 |
||
511 |
Goal "((z1::hypreal) * z2) * z3 = z1 * (z2 * z3)"; |
|
512 |
by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1); |
|
513 |
by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1); |
|
514 |
by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1); |
|
515 |
by (asm_simp_tac (simpset() addsimps [hypreal_mult,real_mult_assoc]) 1); |
|
516 |
qed "hypreal_mult_assoc"; |
|
517 |
||
9432 | 518 |
qed_goal "hypreal_mult_left_commute" (the_context ()) |
7218 | 519 |
"(z1::hypreal) * (z2 * z3) = z2 * (z1 * z3)" |
10648 | 520 |
(fn _ => [rtac (hypreal_mult_commute RS trans) 1, |
521 |
rtac (hypreal_mult_assoc RS trans) 1, |
|
7218 | 522 |
rtac (hypreal_mult_commute RS arg_cong) 1]); |
523 |
||
524 |
(* hypreal multiplication is an AC operator *) |
|
9108 | 525 |
bind_thms ("hypreal_mult_ac", [hypreal_mult_assoc, hypreal_mult_commute, |
526 |
hypreal_mult_left_commute]); |
|
7218 | 527 |
|
528 |
Goalw [hypreal_one_def] "1hr * z = z"; |
|
529 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
|
530 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult]) 1); |
|
531 |
qed "hypreal_mult_1"; |
|
532 |
||
533 |
Goal "z * 1hr = z"; |
|
534 |
by (simp_tac (simpset() addsimps [hypreal_mult_commute, |
|
535 |
hypreal_mult_1]) 1); |
|
536 |
qed "hypreal_mult_1_right"; |
|
537 |
||
9055 | 538 |
Goalw [hypreal_zero_def] "0 * z = (0::hypreal)"; |
7218 | 539 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
540 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult,real_mult_0]) 1); |
|
541 |
qed "hypreal_mult_0"; |
|
542 |
||
9055 | 543 |
Goal "z * 0 = (0::hypreal)"; |
10677 | 544 |
by (simp_tac (simpset() addsimps [hypreal_mult_commute, hypreal_mult_0]) 1); |
7218 | 545 |
qed "hypreal_mult_0_right"; |
546 |
||
547 |
Addsimps [hypreal_mult_0,hypreal_mult_0_right]; |
|
548 |
||
549 |
Goal "-(x * y) = -x * (y::hypreal)"; |
|
550 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
551 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
9043
ca761fe227d8
First round of changes, towards installation of simprocs
paulson
parents:
9013
diff
changeset
|
552 |
by (auto_tac (claset(), |
ca761fe227d8
First round of changes, towards installation of simprocs
paulson
parents:
9013
diff
changeset
|
553 |
simpset() addsimps [hypreal_minus, hypreal_mult] |
ca761fe227d8
First round of changes, towards installation of simprocs
paulson
parents:
9013
diff
changeset
|
554 |
@ real_mult_ac @ real_add_ac)); |
7218 | 555 |
qed "hypreal_minus_mult_eq1"; |
556 |
||
557 |
Goal "-(x * y) = (x::hypreal) * -y"; |
|
558 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
559 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
10677 | 560 |
by (auto_tac (claset(), simpset() addsimps [hypreal_minus, hypreal_mult] |
9043
ca761fe227d8
First round of changes, towards installation of simprocs
paulson
parents:
9013
diff
changeset
|
561 |
@ real_mult_ac @ real_add_ac)); |
7218 | 562 |
qed "hypreal_minus_mult_eq2"; |
563 |
||
9055 | 564 |
(*Pull negations out*) |
565 |
Addsimps [hypreal_minus_mult_eq2 RS sym, hypreal_minus_mult_eq1 RS sym]; |
|
7218 | 566 |
|
567 |
Goal "-x*y = (x::hypreal)*-y"; |
|
9055 | 568 |
by Auto_tac; |
7218 | 569 |
qed "hypreal_minus_mult_commute"; |
570 |
||
571 |
(*----------------------------------------------------------------------------- |
|
572 |
A few more theorems |
|
573 |
----------------------------------------------------------------------------*) |
|
574 |
Goal "(z::hypreal) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"; |
|
575 |
by (asm_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
|
576 |
qed "hypreal_add_assoc_cong"; |
|
577 |
||
578 |
Goal "(z::hypreal) + (v + w) = v + (z + w)"; |
|
579 |
by (REPEAT (ares_tac [hypreal_add_commute RS hypreal_add_assoc_cong] 1)); |
|
580 |
qed "hypreal_add_assoc_swap"; |
|
581 |
||
582 |
Goal "((z1::hypreal) + z2) * w = (z1 * w) + (z2 * w)"; |
|
583 |
by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1); |
|
584 |
by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1); |
|
585 |
by (res_inst_tac [("z","w")] eq_Abs_hypreal 1); |
|
586 |
by (asm_simp_tac (simpset() addsimps [hypreal_mult,hypreal_add, |
|
587 |
real_add_mult_distrib]) 1); |
|
588 |
qed "hypreal_add_mult_distrib"; |
|
589 |
||
590 |
val hypreal_mult_commute'= read_instantiate [("z","w")] hypreal_mult_commute; |
|
591 |
||
592 |
Goal "(w::hypreal) * (z1 + z2) = (w * z1) + (w * z2)"; |
|
593 |
by (simp_tac (simpset() addsimps [hypreal_mult_commute',hypreal_add_mult_distrib]) 1); |
|
594 |
qed "hypreal_add_mult_distrib2"; |
|
595 |
||
9108 | 596 |
bind_thms ("hypreal_mult_simps", [hypreal_mult_1, hypreal_mult_1_right]); |
7218 | 597 |
Addsimps hypreal_mult_simps; |
598 |
||
10043 | 599 |
(* 07/00 *) |
600 |
||
601 |
Goalw [hypreal_diff_def] "((z1::hypreal) - z2) * w = (z1 * w) - (z2 * w)"; |
|
602 |
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1); |
|
603 |
qed "hypreal_diff_mult_distrib"; |
|
604 |
||
605 |
Goal "(w::hypreal) * (z1 - z2) = (w * z1) - (w * z2)"; |
|
606 |
by (simp_tac (simpset() addsimps [hypreal_mult_commute', |
|
607 |
hypreal_diff_mult_distrib]) 1); |
|
608 |
qed "hypreal_diff_mult_distrib2"; |
|
609 |
||
7218 | 610 |
(*** one and zero are distinct ***) |
9055 | 611 |
Goalw [hypreal_zero_def,hypreal_one_def] "0 ~= 1hr"; |
10677 | 612 |
by (auto_tac (claset(), simpset() addsimps [real_zero_not_eq_one])); |
7218 | 613 |
qed "hypreal_zero_not_eq_one"; |
614 |
||
10677 | 615 |
|
616 |
(**** multiplicative inverse on hypreal ****) |
|
617 |
||
618 |
Goalw [congruent_def] |
|
619 |
"congruent hyprel (%X. hyprel^^{%n. if X n = #0 then #0 else inverse(X n)})"; |
|
620 |
by (Auto_tac THEN Ultra_tac 1); |
|
621 |
qed "hypreal_inverse_congruent"; |
|
622 |
||
623 |
Goalw [hypreal_inverse_def] |
|
624 |
"inverse (Abs_hypreal(hyprel^^{%n. X n})) = \ |
|
625 |
\ Abs_hypreal(hyprel ^^ {%n. if X n = #0 then #0 else inverse(X n)})"; |
|
626 |
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); |
|
627 |
by (simp_tac (simpset() addsimps |
|
628 |
[hyprel_in_hypreal RS Abs_hypreal_inverse, |
|
629 |
[equiv_hyprel, hypreal_inverse_congruent] MRS UN_equiv_class]) 1); |
|
630 |
qed "hypreal_inverse"; |
|
631 |
||
632 |
Goal "inverse 0 = (0::hypreal)"; |
|
633 |
by (simp_tac (simpset() addsimps [hypreal_inverse, hypreal_zero_def]) 1); |
|
634 |
qed "HYPREAL_INVERSE_ZERO"; |
|
635 |
||
636 |
Goal "a / (0::hypreal) = 0"; |
|
637 |
by (simp_tac |
|
638 |
(simpset() addsimps [hypreal_divide_def, HYPREAL_INVERSE_ZERO]) 1); |
|
639 |
qed "HYPREAL_DIVISION_BY_ZERO"; (*NOT for adding to default simpset*) |
|
640 |
||
641 |
fun hypreal_div_undefined_case_tac s i = |
|
642 |
case_tac s i THEN |
|
643 |
asm_simp_tac |
|
644 |
(simpset() addsimps [HYPREAL_DIVISION_BY_ZERO, HYPREAL_INVERSE_ZERO]) i; |
|
645 |
||
646 |
Goal "inverse (inverse (z::hypreal)) = z"; |
|
647 |
by (hypreal_div_undefined_case_tac "z=0" 1); |
|
648 |
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); |
|
649 |
by (asm_full_simp_tac (simpset() addsimps |
|
650 |
[hypreal_inverse, hypreal_zero_def]) 1); |
|
651 |
qed "hypreal_inverse_inverse"; |
|
652 |
Addsimps [hypreal_inverse_inverse]; |
|
653 |
||
654 |
Goalw [hypreal_one_def] "inverse(1hr) = 1hr"; |
|
655 |
by (full_simp_tac (simpset() addsimps [hypreal_inverse, |
|
656 |
real_zero_not_eq_one RS not_sym]) 1); |
|
657 |
qed "hypreal_inverse_1"; |
|
658 |
Addsimps [hypreal_inverse_1]; |
|
659 |
||
660 |
||
7218 | 661 |
(*** existence of inverse ***) |
10677 | 662 |
|
7218 | 663 |
Goalw [hypreal_one_def,hypreal_zero_def] |
10677 | 664 |
"x ~= 0 ==> x*inverse(x) = 1hr"; |
7218 | 665 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
666 |
by (rotate_tac 1 1); |
|
10677 | 667 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1); |
7218 | 668 |
by (dtac FreeUltrafilterNat_Compl_mem 1); |
669 |
by (blast_tac (claset() addSIs [real_mult_inv_right, |
|
670 |
FreeUltrafilterNat_subset]) 1); |
|
10607 | 671 |
qed "hypreal_mult_inverse"; |
7218 | 672 |
|
10607 | 673 |
Goal "x ~= 0 ==> inverse(x)*x = 1hr"; |
674 |
by (asm_simp_tac (simpset() addsimps [hypreal_mult_inverse, |
|
9055 | 675 |
hypreal_mult_commute]) 1); |
10607 | 676 |
qed "hypreal_mult_inverse_left"; |
7218 | 677 |
|
9055 | 678 |
Goal "(c::hypreal) ~= 0 ==> (c*a=c*b) = (a=b)"; |
7218 | 679 |
by Auto_tac; |
10607 | 680 |
by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1); |
681 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac) 1); |
|
7218 | 682 |
qed "hypreal_mult_left_cancel"; |
683 |
||
9055 | 684 |
Goal "(c::hypreal) ~= 0 ==> (a*c=b*c) = (a=b)"; |
7218 | 685 |
by (Step_tac 1); |
10607 | 686 |
by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1); |
687 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac) 1); |
|
7218 | 688 |
qed "hypreal_mult_right_cancel"; |
689 |
||
10607 | 690 |
Goalw [hypreal_zero_def] "x ~= 0 ==> inverse (x::hypreal) ~= 0"; |
7218 | 691 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
10677 | 692 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1); |
10607 | 693 |
qed "hypreal_inverse_not_zero"; |
7218 | 694 |
|
10607 | 695 |
Addsimps [hypreal_mult_inverse,hypreal_mult_inverse_left]; |
7218 | 696 |
|
9055 | 697 |
Goal "[| x ~= 0; y ~= 0 |] ==> x * y ~= (0::hypreal)"; |
7218 | 698 |
by (Step_tac 1); |
10607 | 699 |
by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1); |
7218 | 700 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1); |
701 |
qed "hypreal_mult_not_0"; |
|
702 |
||
703 |
bind_thm ("hypreal_mult_not_0E",hypreal_mult_not_0 RS notE); |
|
704 |
||
10043 | 705 |
Goal "x*y = (0::hypreal) ==> x = 0 | y = 0"; |
10677 | 706 |
by (auto_tac (claset() addIs [ccontr] addDs [hypreal_mult_not_0], |
707 |
simpset())); |
|
10043 | 708 |
qed "hypreal_mult_zero_disj"; |
709 |
||
9055 | 710 |
Goal "x ~= 0 ==> x * x ~= (0::hypreal)"; |
7218 | 711 |
by (blast_tac (claset() addDs [hypreal_mult_not_0]) 1); |
712 |
qed "hypreal_mult_self_not_zero"; |
|
713 |
||
10607 | 714 |
Goal "[| x ~= 0; y ~= 0 |] ==> inverse(x*y) = inverse(x)*inverse(y::hypreal)"; |
7218 | 715 |
by (res_inst_tac [("c1","x")] (hypreal_mult_left_cancel RS iffD1) 1); |
10677 | 716 |
by (auto_tac (claset(), |
717 |
simpset() addsimps [hypreal_mult_assoc RS sym, hypreal_mult_not_0])); |
|
7218 | 718 |
by (res_inst_tac [("c1","y")] (hypreal_mult_right_cancel RS iffD1) 1); |
10677 | 719 |
by (auto_tac (claset(), |
720 |
simpset() addsimps [hypreal_mult_not_0] @ hypreal_mult_ac)); |
|
721 |
by (auto_tac (claset(), |
|
722 |
simpset() addsimps [hypreal_mult_assoc RS sym, hypreal_mult_not_0])); |
|
10607 | 723 |
qed "inverse_mult_eq"; |
7218 | 724 |
|
10677 | 725 |
Goal "inverse(-x) = -inverse(x::hypreal)"; |
726 |
by (hypreal_div_undefined_case_tac "x=0" 1); |
|
9055 | 727 |
by (rtac (hypreal_mult_right_cancel RS iffD1) 1); |
10607 | 728 |
by (stac hypreal_mult_inverse_left 2); |
7218 | 729 |
by Auto_tac; |
10607 | 730 |
qed "hypreal_minus_inverse"; |
7218 | 731 |
|
10677 | 732 |
Goal "inverse(x*y) = inverse(x)*inverse(y::hypreal)"; |
733 |
by (hypreal_div_undefined_case_tac "x=0" 1); |
|
734 |
by (hypreal_div_undefined_case_tac "y=0" 1); |
|
7218 | 735 |
by (forw_inst_tac [("y","y")] hypreal_mult_not_0 1 THEN assume_tac 1); |
736 |
by (res_inst_tac [("c1","x")] (hypreal_mult_left_cancel RS iffD1) 1); |
|
10677 | 737 |
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_assoc RS sym])); |
7218 | 738 |
by (res_inst_tac [("c1","y")] (hypreal_mult_left_cancel RS iffD1) 1); |
10677 | 739 |
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_commute])); |
7218 | 740 |
by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1); |
10607 | 741 |
qed "hypreal_inverse_distrib"; |
7218 | 742 |
|
743 |
(*------------------------------------------------------------------ |
|
744 |
Theorems for ordering |
|
745 |
------------------------------------------------------------------*) |
|
746 |
||
747 |
(* prove introduction and elimination rules for hypreal_less *) |
|
748 |
||
749 |
Goalw [hypreal_less_def] |
|
750 |
"P < (Q::hypreal) = (EX X Y. X : Rep_hypreal(P) & \ |
|
751 |
\ Y : Rep_hypreal(Q) & \ |
|
752 |
\ {n. X n < Y n} : FreeUltrafilterNat)"; |
|
753 |
by (Fast_tac 1); |
|
754 |
qed "hypreal_less_iff"; |
|
755 |
||
756 |
Goalw [hypreal_less_def] |
|
757 |
"[| {n. X n < Y n} : FreeUltrafilterNat; \ |
|
758 |
\ X : Rep_hypreal(P); \ |
|
759 |
\ Y : Rep_hypreal(Q) |] ==> P < (Q::hypreal)"; |
|
760 |
by (Fast_tac 1); |
|
761 |
qed "hypreal_lessI"; |
|
762 |
||
763 |
||
764 |
Goalw [hypreal_less_def] |
|
765 |
"!! R1. [| R1 < (R2::hypreal); \ |
|
766 |
\ !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \ |
|
767 |
\ !!X. X : Rep_hypreal(R1) ==> P; \ |
|
768 |
\ !!Y. Y : Rep_hypreal(R2) ==> P |] \ |
|
769 |
\ ==> P"; |
|
770 |
by Auto_tac; |
|
771 |
qed "hypreal_lessE"; |
|
772 |
||
773 |
Goalw [hypreal_less_def] |
|
774 |
"R1 < (R2::hypreal) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \ |
|
775 |
\ X : Rep_hypreal(R1) & \ |
|
776 |
\ Y : Rep_hypreal(R2))"; |
|
777 |
by (Fast_tac 1); |
|
778 |
qed "hypreal_lessD"; |
|
779 |
||
780 |
Goal "~ (R::hypreal) < R"; |
|
781 |
by (res_inst_tac [("z","R")] eq_Abs_hypreal 1); |
|
10677 | 782 |
by (auto_tac (claset(), simpset() addsimps [hypreal_less_def])); |
7218 | 783 |
by (Ultra_tac 1); |
784 |
qed "hypreal_less_not_refl"; |
|
785 |
||
786 |
(*** y < y ==> P ***) |
|
787 |
bind_thm("hypreal_less_irrefl",hypreal_less_not_refl RS notE); |
|
10043 | 788 |
AddSEs [hypreal_less_irrefl]; |
7218 | 789 |
|
790 |
Goal "!!(x::hypreal). x < y ==> x ~= y"; |
|
10677 | 791 |
by (auto_tac (claset(), simpset() addsimps [hypreal_less_not_refl])); |
7218 | 792 |
qed "hypreal_not_refl2"; |
793 |
||
794 |
Goal "!!(R1::hypreal). [| R1 < R2; R2 < R3 |] ==> R1 < R3"; |
|
795 |
by (res_inst_tac [("z","R1")] eq_Abs_hypreal 1); |
|
796 |
by (res_inst_tac [("z","R2")] eq_Abs_hypreal 1); |
|
797 |
by (res_inst_tac [("z","R3")] eq_Abs_hypreal 1); |
|
10677 | 798 |
by (auto_tac (claset() addSIs [exI], simpset() addsimps [hypreal_less_def])); |
799 |
by (ultra_tac (claset() addIs [real_less_trans], simpset()) 1); |
|
7218 | 800 |
qed "hypreal_less_trans"; |
801 |
||
802 |
Goal "!! (R1::hypreal). [| R1 < R2; R2 < R1 |] ==> P"; |
|
803 |
by (dtac hypreal_less_trans 1 THEN assume_tac 1); |
|
804 |
by (asm_full_simp_tac (simpset() addsimps |
|
805 |
[hypreal_less_not_refl]) 1); |
|
806 |
qed "hypreal_less_asym"; |
|
807 |
||
10043 | 808 |
(*------------------------------------------------------- |
7218 | 809 |
TODO: The following theorem should have been proved |
810 |
first and then used througout the proofs as it probably |
|
811 |
makes many of them more straightforward. |
|
812 |
-------------------------------------------------------*) |
|
813 |
Goalw [hypreal_less_def] |
|
814 |
"(Abs_hypreal(hyprel^^{%n. X n}) < \ |
|
815 |
\ Abs_hypreal(hyprel^^{%n. Y n})) = \ |
|
816 |
\ ({n. X n < Y n} : FreeUltrafilterNat)"; |
|
10677 | 817 |
by (auto_tac (claset() addSIs [lemma_hyprel_refl], simpset())); |
7218 | 818 |
by (Ultra_tac 1); |
819 |
qed "hypreal_less"; |
|
820 |
||
821 |
(*--------------------------------------------------------------------------------- |
|
822 |
Hyperreals as a linearly ordered field |
|
823 |
---------------------------------------------------------------------------------*) |
|
10607 | 824 |
(*** sum order |
7218 | 825 |
Goalw [hypreal_zero_def] |
9055 | 826 |
"[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y"; |
7218 | 827 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
828 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
829 |
by (auto_tac (claset(),simpset() addsimps |
|
830 |
[hypreal_less_def,hypreal_add])); |
|
831 |
by (auto_tac (claset() addSIs [exI],simpset() addsimps |
|
832 |
[hypreal_less_def,hypreal_add])); |
|
833 |
by (ultra_tac (claset() addIs [real_add_order],simpset()) 1); |
|
834 |
qed "hypreal_add_order"; |
|
10607 | 835 |
***) |
7218 | 836 |
|
10607 | 837 |
(*** mult order |
7218 | 838 |
Goalw [hypreal_zero_def] |
9055 | 839 |
"[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y"; |
7218 | 840 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
841 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
842 |
by (auto_tac (claset() addSIs [exI],simpset() addsimps |
|
843 |
[hypreal_less_def,hypreal_mult])); |
|
9432 | 844 |
by (ultra_tac (claset() addIs [rename_numerals real_mult_order], |
9071 | 845 |
simpset()) 1); |
7218 | 846 |
qed "hypreal_mult_order"; |
10043 | 847 |
****) |
848 |
||
7218 | 849 |
|
850 |
(*--------------------------------------------------------------------------------- |
|
851 |
Trichotomy of the hyperreals |
|
852 |
--------------------------------------------------------------------------------*) |
|
853 |
||
9055 | 854 |
Goalw [hyprel_def] "EX x. x: hyprel ^^ {%n. #0}"; |
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8856
diff
changeset
|
855 |
by (res_inst_tac [("x","%n. #0")] exI 1); |
7218 | 856 |
by (Step_tac 1); |
10677 | 857 |
by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset())); |
7218 | 858 |
qed "lemma_hyprel_0r_mem"; |
859 |
||
9055 | 860 |
Goalw [hypreal_zero_def]"0 < x | x = 0 | x < (0::hypreal)"; |
7218 | 861 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
862 |
by (auto_tac (claset(),simpset() addsimps [hypreal_less_def])); |
|
863 |
by (cut_facts_tac [lemma_hyprel_0r_mem] 1 THEN etac exE 1); |
|
864 |
by (dres_inst_tac [("x","xa")] spec 1); |
|
865 |
by (dres_inst_tac [("x","x")] spec 1); |
|
866 |
by (cut_inst_tac [("x","x")] lemma_hyprel_refl 1); |
|
867 |
by Auto_tac; |
|
868 |
by (dres_inst_tac [("x","x")] spec 1); |
|
869 |
by (dres_inst_tac [("x","xa")] spec 1); |
|
870 |
by Auto_tac; |
|
871 |
by (Ultra_tac 1); |
|
872 |
by (auto_tac (claset() addIs [real_linear_less2],simpset())); |
|
873 |
qed "hypreal_trichotomy"; |
|
874 |
||
9055 | 875 |
val prems = Goal "[| (0::hypreal) < x ==> P; \ |
876 |
\ x = 0 ==> P; \ |
|
877 |
\ x < 0 ==> P |] ==> P"; |
|
7218 | 878 |
by (cut_inst_tac [("x","x")] hypreal_trichotomy 1); |
879 |
by (REPEAT (eresolve_tac (disjE::prems) 1)); |
|
880 |
qed "hypreal_trichotomyE"; |
|
881 |
||
882 |
(*---------------------------------------------------------------------------- |
|
883 |
More properties of < |
|
884 |
----------------------------------------------------------------------------*) |
|
885 |
||
9055 | 886 |
Goal "((x::hypreal) < y) = (0 < y + -x)"; |
10043 | 887 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
888 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
889 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add, |
|
890 |
hypreal_zero_def,hypreal_minus,hypreal_less])); |
|
891 |
by (ALLGOALS(Ultra_tac)); |
|
7218 | 892 |
qed "hypreal_less_minus_iff"; |
893 |
||
10043 | 894 |
Goal "((x::hypreal) < y) = (x + -y < 0)"; |
895 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
|
896 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
897 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add, |
|
898 |
hypreal_zero_def,hypreal_minus,hypreal_less])); |
|
899 |
by (ALLGOALS(Ultra_tac)); |
|
7218 | 900 |
qed "hypreal_less_minus_iff2"; |
901 |
||
9055 | 902 |
Goal "((x::hypreal) = y) = (0 = x + - y)"; |
7218 | 903 |
by Auto_tac; |
904 |
by (res_inst_tac [("x1","-y")] (hypreal_add_right_cancel RS iffD1) 1); |
|
905 |
by Auto_tac; |
|
906 |
qed "hypreal_eq_minus_iff"; |
|
907 |
||
9055 | 908 |
Goal "((x::hypreal) = y) = (0 = y + - x)"; |
7218 | 909 |
by Auto_tac; |
910 |
by (res_inst_tac [("x1","-x")] (hypreal_add_right_cancel RS iffD1) 1); |
|
911 |
by Auto_tac; |
|
912 |
qed "hypreal_eq_minus_iff2"; |
|
913 |
||
10043 | 914 |
(* 07/00 *) |
915 |
Goal "(0::hypreal) - x = -x"; |
|
916 |
by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1); |
|
917 |
qed "hypreal_diff_zero"; |
|
918 |
||
919 |
Goal "x - (0::hypreal) = x"; |
|
920 |
by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1); |
|
921 |
qed "hypreal_diff_zero_right"; |
|
922 |
||
923 |
Goal "x - x = (0::hypreal)"; |
|
924 |
by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1); |
|
925 |
qed "hypreal_diff_self"; |
|
926 |
||
927 |
Addsimps [hypreal_diff_zero, hypreal_diff_zero_right, hypreal_diff_self]; |
|
928 |
||
7218 | 929 |
Goal "(x = y + z) = (x + -z = (y::hypreal))"; |
930 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc])); |
|
931 |
qed "hypreal_eq_minus_iff3"; |
|
932 |
||
933 |
Goal "(x = z + y) = (x + -z = (y::hypreal))"; |
|
934 |
by (auto_tac (claset(),simpset() addsimps hypreal_add_ac)); |
|
935 |
qed "hypreal_eq_minus_iff4"; |
|
936 |
||
9055 | 937 |
Goal "(x ~= a) = (x + -a ~= (0::hypreal))"; |
7218 | 938 |
by (auto_tac (claset() addDs [sym RS |
939 |
(hypreal_eq_minus_iff RS iffD2)],simpset())); |
|
940 |
qed "hypreal_not_eq_minus_iff"; |
|
941 |
||
942 |
(*** linearity ***) |
|
943 |
Goal "(x::hypreal) < y | x = y | y < x"; |
|
7322 | 944 |
by (stac hypreal_eq_minus_iff2 1); |
7218 | 945 |
by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1); |
946 |
by (res_inst_tac [("x1","y")] (hypreal_less_minus_iff2 RS ssubst) 1); |
|
947 |
by (rtac hypreal_trichotomyE 1); |
|
948 |
by Auto_tac; |
|
949 |
qed "hypreal_linear"; |
|
950 |
||
10043 | 951 |
Goal "((w::hypreal) ~= z) = (w<z | z<w)"; |
952 |
by (cut_facts_tac [hypreal_linear] 1); |
|
953 |
by (Blast_tac 1); |
|
954 |
qed "hypreal_neq_iff"; |
|
955 |
||
7218 | 956 |
Goal "!!(x::hypreal). [| x < y ==> P; x = y ==> P; \ |
957 |
\ y < x ==> P |] ==> P"; |
|
958 |
by (cut_inst_tac [("x","x"),("y","y")] hypreal_linear 1); |
|
959 |
by Auto_tac; |
|
960 |
qed "hypreal_linear_less2"; |
|
961 |
||
962 |
(*------------------------------------------------------------------------------ |
|
963 |
Properties of <= |
|
964 |
------------------------------------------------------------------------------*) |
|
965 |
(*------ hypreal le iff reals le a.e ------*) |
|
966 |
||
967 |
Goalw [hypreal_le_def,real_le_def] |
|
968 |
"(Abs_hypreal(hyprel^^{%n. X n}) <= \ |
|
969 |
\ Abs_hypreal(hyprel^^{%n. Y n})) = \ |
|
970 |
\ ({n. X n <= Y n} : FreeUltrafilterNat)"; |
|
971 |
by (auto_tac (claset(),simpset() addsimps [hypreal_less])); |
|
972 |
by (ALLGOALS(Ultra_tac)); |
|
973 |
qed "hypreal_le"; |
|
974 |
||
975 |
(*---------------------------------------------------------*) |
|
976 |
(*---------------------------------------------------------*) |
|
977 |
Goalw [hypreal_le_def] |
|
978 |
"~(w < z) ==> z <= (w::hypreal)"; |
|
979 |
by (assume_tac 1); |
|
980 |
qed "hypreal_leI"; |
|
981 |
||
982 |
Goalw [hypreal_le_def] |
|
983 |
"z<=w ==> ~(w<(z::hypreal))"; |
|
984 |
by (assume_tac 1); |
|
985 |
qed "hypreal_leD"; |
|
986 |
||
9108 | 987 |
bind_thm ("hypreal_leE", make_elim hypreal_leD); |
7218 | 988 |
|
989 |
Goal "(~(w < z)) = (z <= (w::hypreal))"; |
|
990 |
by (fast_tac (claset() addSIs [hypreal_leI,hypreal_leD]) 1); |
|
991 |
qed "hypreal_less_le_iff"; |
|
992 |
||
993 |
Goalw [hypreal_le_def] "~ z <= w ==> w<(z::hypreal)"; |
|
994 |
by (Fast_tac 1); |
|
995 |
qed "not_hypreal_leE"; |
|
996 |
||
997 |
Goalw [hypreal_le_def] "z < w ==> z <= (w::hypreal)"; |
|
998 |
by (fast_tac (claset() addEs [hypreal_less_asym]) 1); |
|
999 |
qed "hypreal_less_imp_le"; |
|
1000 |
||
1001 |
Goalw [hypreal_le_def] "!!(x::hypreal). x <= y ==> x < y | x = y"; |
|
1002 |
by (cut_facts_tac [hypreal_linear] 1); |
|
1003 |
by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1); |
|
1004 |
qed "hypreal_le_imp_less_or_eq"; |
|
1005 |
||
1006 |
Goalw [hypreal_le_def] "z<w | z=w ==> z <=(w::hypreal)"; |
|
1007 |
by (cut_facts_tac [hypreal_linear] 1); |
|
1008 |
by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1); |
|
1009 |
qed "hypreal_less_or_eq_imp_le"; |
|
1010 |
||
1011 |
Goal "(x <= (y::hypreal)) = (x < y | x=y)"; |
|
1012 |
by (REPEAT(ares_tac [iffI, hypreal_less_or_eq_imp_le, hypreal_le_imp_less_or_eq] 1)); |
|
1013 |
qed "hypreal_le_eq_less_or_eq"; |
|
10043 | 1014 |
val hypreal_le_less = hypreal_le_eq_less_or_eq; |
7218 | 1015 |
|
1016 |
Goal "w <= (w::hypreal)"; |
|
1017 |
by (simp_tac (simpset() addsimps [hypreal_le_eq_less_or_eq]) 1); |
|
1018 |
qed "hypreal_le_refl"; |
|
1019 |
Addsimps [hypreal_le_refl]; |
|
1020 |
||
10043 | 1021 |
(* Axiom 'linorder_linear' of class 'linorder': *) |
1022 |
Goal "(z::hypreal) <= w | w <= z"; |
|
1023 |
by (simp_tac (simpset() addsimps [hypreal_le_less]) 1); |
|
1024 |
by (cut_facts_tac [hypreal_linear] 1); |
|
1025 |
by (Blast_tac 1); |
|
1026 |
qed "hypreal_le_linear"; |
|
1027 |
||
7218 | 1028 |
Goal "[| i <= j; j < k |] ==> i < (k::hypreal)"; |
1029 |
by (dtac hypreal_le_imp_less_or_eq 1); |
|
1030 |
by (fast_tac (claset() addIs [hypreal_less_trans]) 1); |
|
1031 |
qed "hypreal_le_less_trans"; |
|
1032 |
||
1033 |
Goal "!! (i::hypreal). [| i < j; j <= k |] ==> i < k"; |
|
1034 |
by (dtac hypreal_le_imp_less_or_eq 1); |
|
1035 |
by (fast_tac (claset() addIs [hypreal_less_trans]) 1); |
|
1036 |
qed "hypreal_less_le_trans"; |
|
1037 |
||
1038 |
Goal "[| i <= j; j <= k |] ==> i <= (k::hypreal)"; |
|
1039 |
by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq, |
|
1040 |
rtac hypreal_less_or_eq_imp_le, fast_tac (claset() addIs [hypreal_less_trans])]); |
|
1041 |
qed "hypreal_le_trans"; |
|
1042 |
||
1043 |
Goal "[| z <= w; w <= z |] ==> z = (w::hypreal)"; |
|
1044 |
by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq, |
|
1045 |
fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym])]); |
|
1046 |
qed "hypreal_le_anti_sym"; |
|
1047 |
||
1048 |
Goal "[| ~ y < x; y ~= x |] ==> x < (y::hypreal)"; |
|
1049 |
by (rtac not_hypreal_leE 1); |
|
1050 |
by (fast_tac (claset() addDs [hypreal_le_imp_less_or_eq]) 1); |
|
1051 |
qed "not_less_not_eq_hypreal_less"; |
|
1052 |
||
10043 | 1053 |
(* Axiom 'order_less_le' of class 'order': *) |
1054 |
Goal "(w::hypreal) < z = (w <= z & w ~= z)"; |
|
1055 |
by (simp_tac (simpset() addsimps [hypreal_le_def, hypreal_neq_iff]) 1); |
|
1056 |
by (blast_tac (claset() addIs [hypreal_less_asym]) 1); |
|
1057 |
qed "hypreal_less_le"; |
|
1058 |
||
9055 | 1059 |
Goal "(0 < -R) = (R < (0::hypreal))"; |
10043 | 1060 |
by (res_inst_tac [("z","R")] eq_Abs_hypreal 1); |
1061 |
by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def, |
|
1062 |
hypreal_less,hypreal_minus])); |
|
7218 | 1063 |
qed "hypreal_minus_zero_less_iff"; |
10043 | 1064 |
Addsimps [hypreal_minus_zero_less_iff]; |
7218 | 1065 |
|
9055 | 1066 |
Goal "(-R < 0) = ((0::hypreal) < R)"; |
10043 | 1067 |
by (res_inst_tac [("z","R")] eq_Abs_hypreal 1); |
1068 |
by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def, |
|
1069 |
hypreal_less,hypreal_minus])); |
|
1070 |
by (ALLGOALS(Ultra_tac)); |
|
7218 | 1071 |
qed "hypreal_minus_zero_less_iff2"; |
1072 |
||
10043 | 1073 |
Goalw [hypreal_le_def] "((0::hypreal) <= -r) = (r <= (0::hypreal))"; |
1074 |
by (simp_tac (simpset() addsimps |
|
1075 |
[hypreal_minus_zero_less_iff2]) 1); |
|
1076 |
qed "hypreal_minus_zero_le_iff"; |
|
7218 | 1077 |
|
1078 |
(*---------------------------------------------------------- |
|
1079 |
hypreal_of_real preserves field and order properties |
|
1080 |
-----------------------------------------------------------*) |
|
1081 |
Goalw [hypreal_of_real_def] |
|
9071 | 1082 |
"hypreal_of_real (z1 + z2) = \ |
1083 |
\ hypreal_of_real z1 + hypreal_of_real z2"; |
|
7218 | 1084 |
by (asm_simp_tac (simpset() addsimps [hypreal_add, |
1085 |
hypreal_add_mult_distrib]) 1); |
|
1086 |
qed "hypreal_of_real_add"; |
|
1087 |
||
1088 |
Goalw [hypreal_of_real_def] |
|
9071 | 1089 |
"hypreal_of_real (z1 * z2) = hypreal_of_real z1 * hypreal_of_real z2"; |
7218 | 1090 |
by (full_simp_tac (simpset() addsimps [hypreal_mult, |
1091 |
hypreal_add_mult_distrib2]) 1); |
|
1092 |
qed "hypreal_of_real_mult"; |
|
1093 |
||
1094 |
Goalw [hypreal_less_def,hypreal_of_real_def] |
|
1095 |
"(z1 < z2) = (hypreal_of_real z1 < hypreal_of_real z2)"; |
|
1096 |
by Auto_tac; |
|
1097 |
by (res_inst_tac [("x","%n. z1")] exI 1); |
|
1098 |
by (Step_tac 1); |
|
1099 |
by (res_inst_tac [("x","%n. z2")] exI 2); |
|
1100 |
by Auto_tac; |
|
1101 |
by (rtac FreeUltrafilterNat_P 1); |
|
1102 |
by (Ultra_tac 1); |
|
1103 |
qed "hypreal_of_real_less_iff"; |
|
1104 |
||
1105 |
Addsimps [hypreal_of_real_less_iff RS sym]; |
|
1106 |
||
1107 |
Goalw [hypreal_le_def,real_le_def] |
|
1108 |
"(z1 <= z2) = (hypreal_of_real z1 <= hypreal_of_real z2)"; |
|
1109 |
by Auto_tac; |
|
1110 |
qed "hypreal_of_real_le_iff"; |
|
1111 |
||
1112 |
Goalw [hypreal_of_real_def] "hypreal_of_real (-r) = - hypreal_of_real r"; |
|
1113 |
by (auto_tac (claset(),simpset() addsimps [hypreal_minus])); |
|
1114 |
qed "hypreal_of_real_minus"; |
|
1115 |
||
10677 | 1116 |
(*DON'T insert this or the next one as default simprules. |
1117 |
They are used in both orientations and anyway aren't the ones we finally |
|
1118 |
need, which would use binary literals.*) |
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
8856
diff
changeset
|
1119 |
Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real #1 = 1hr"; |
7218 | 1120 |
by (Step_tac 1); |
1121 |
qed "hypreal_of_real_one"; |
|
1122 |
||
10677 | 1123 |
Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real #0 = 0"; |
7218 | 1124 |
by (Step_tac 1); |
1125 |
qed "hypreal_of_real_zero"; |
|
1126 |
||
10677 | 1127 |
Goal "(hypreal_of_real r = 0) = (r = #0)"; |
7218 | 1128 |
by (auto_tac (claset() addIs [FreeUltrafilterNat_P], |
1129 |
simpset() addsimps [hypreal_of_real_def, |
|
10677 | 1130 |
hypreal_zero_def,FreeUltrafilterNat_Nat_set])); |
7218 | 1131 |
qed "hypreal_of_real_zero_iff"; |
1132 |
||
10677 | 1133 |
(*FIXME: delete*) |
1134 |
Goal "(hypreal_of_real r ~= 0) = (r ~= #0)"; |
|
7218 | 1135 |
by (full_simp_tac (simpset() addsimps [hypreal_of_real_zero_iff]) 1); |
1136 |
qed "hypreal_of_real_not_zero_iff"; |
|
1137 |
||
10677 | 1138 |
Goal "inverse (hypreal_of_real r) = hypreal_of_real (inverse r)"; |
1139 |
by (case_tac "r=#0" 1); |
|
1140 |
by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO, INVERSE_ZERO, |
|
1141 |
HYPREAL_INVERSE_ZERO, hypreal_of_real_zero]) 1); |
|
1142 |
by (res_inst_tac [("c1","hypreal_of_real r")] |
|
1143 |
(hypreal_mult_left_cancel RS iffD1) 1); |
|
7218 | 1144 |
by (etac (hypreal_of_real_not_zero_iff RS iffD2) 1); |
1145 |
by (forward_tac [hypreal_of_real_not_zero_iff RS iffD2] 1); |
|
10677 | 1146 |
by (auto_tac (claset(), |
1147 |
simpset() addsimps [hypreal_of_real_one, hypreal_of_real_mult RS sym])); |
|
10607 | 1148 |
qed "hypreal_of_real_inverse"; |
7218 | 1149 |
|
10690 | 1150 |
Goal "hypreal_of_real (z1 / z2) = hypreal_of_real z1 / hypreal_of_real z2"; |
1151 |
by (simp_tac (simpset() addsimps [hypreal_divide_def, real_divide_def, |
|
1152 |
hypreal_of_real_mult, hypreal_of_real_inverse]) 1); |
|
1153 |
qed "hypreal_of_real_divide"; |
|
1154 |
||
1155 |
||
1156 |
(*** Division lemmas ***) |
|
1157 |
||
1158 |
Goal "(0::hypreal)/x = 0"; |
|
1159 |
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); |
|
1160 |
qed "hypreal_zero_divide"; |
|
1161 |
||
1162 |
Goal "x/1hr = x"; |
|
1163 |
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); |
|
1164 |
qed "hypreal_divide_one"; |
|
1165 |
Addsimps [hypreal_zero_divide, hypreal_divide_one]; |
|
1166 |
||
1167 |
Goal "(x::hypreal) * (y/z) = (x*y)/z"; |
|
1168 |
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1); |
|
1169 |
qed "hypreal_times_divide1_eq"; |
|
1170 |
||
1171 |
Goal "(y/z) * (x::hypreal) = (y*x)/z"; |
|
1172 |
by (simp_tac (simpset() addsimps [hypreal_divide_def]@hypreal_mult_ac) 1); |
|
1173 |
qed "hypreal_times_divide2_eq"; |
|
1174 |
||
1175 |
Addsimps [hypreal_times_divide1_eq, hypreal_times_divide2_eq]; |
|
1176 |
||
1177 |
Goal "(x::hypreal) / (y/z) = (x*z)/y"; |
|
1178 |
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]@ |
|
1179 |
hypreal_mult_ac) 1); |
|
1180 |
qed "hypreal_divide_divide1_eq"; |
|
1181 |
||
1182 |
Goal "((x::hypreal) / y) / z = x/(y*z)"; |
|
1183 |
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib, |
|
1184 |
hypreal_mult_assoc]) 1); |
|
1185 |
qed "hypreal_divide_divide2_eq"; |
|
1186 |
||
1187 |
Addsimps [hypreal_divide_divide1_eq, hypreal_divide_divide2_eq]; |
|
1188 |
||
1189 |
(** As with multiplication, pull minus signs OUT of the / operator **) |
|
1190 |
||
1191 |
Goal "(-x) / (y::hypreal) = - (x/y)"; |
|
1192 |
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); |
|
1193 |
qed "hypreal_minus_divide_eq"; |
|
1194 |
Addsimps [hypreal_minus_divide_eq]; |
|
1195 |
||
1196 |
Goal "(x / -(y::hypreal)) = - (x/y)"; |
|
1197 |
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1); |
|
1198 |
qed "hypreal_divide_minus_eq"; |
|
1199 |
Addsimps [hypreal_divide_minus_eq]; |
|
1200 |
||
1201 |
Goal "(x+y)/(z::hypreal) = x/z + y/z"; |
|
1202 |
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_add_mult_distrib]) 1); |
|
1203 |
qed "hypreal_add_divide_distrib"; |
|
1204 |
||
1205 |
||
7218 | 1206 |
Goal "x+x=x*(1hr+1hr)"; |
1207 |
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2]) 1); |
|
1208 |
qed "hypreal_add_self"; |
|
1209 |
||
10677 | 1210 |
(*FIXME: DELETE (used in Lim.ML) *) |
10607 | 1211 |
Goal "(z::hypreal) ~= 0 ==> x*y = (x*inverse(z))*(z*y)"; |
7218 | 1212 |
by (asm_simp_tac (simpset() addsimps hypreal_mult_ac) 1); |
1213 |
qed "lemma_chain"; |
|
1214 |
||
10607 | 1215 |
Goal "[|(x::hypreal) ~= 0; y ~= 0 |] ==> \ |
1216 |
\ inverse(x) + inverse(y) = (x + y)*inverse(x*y)"; |
|
1217 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse_distrib, |
|
7218 | 1218 |
hypreal_add_mult_distrib,hypreal_mult_assoc RS sym]) 1); |
7322 | 1219 |
by (stac hypreal_mult_assoc 1); |
7218 | 1220 |
by (rtac (hypreal_mult_left_commute RS subst) 1); |
1221 |
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1); |
|
10607 | 1222 |
qed "hypreal_inverse_add"; |
7218 | 1223 |
|
9055 | 1224 |
Goal "x = -x ==> x = (0::hypreal)"; |
10043 | 1225 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
1226 |
by (auto_tac (claset(),simpset() addsimps [hypreal_minus, |
|
1227 |
hypreal_zero_def])); |
|
1228 |
by (Ultra_tac 1); |
|
7218 | 1229 |
qed "hypreal_self_eq_minus_self_zero"; |
1230 |
||
9055 | 1231 |
Goal "(x + x = 0) = (x = (0::hypreal))"; |
10043 | 1232 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
1233 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add, |
|
1234 |
hypreal_zero_def])); |
|
7218 | 1235 |
qed "hypreal_add_self_zero_cancel"; |
1236 |
Addsimps [hypreal_add_self_zero_cancel]; |
|
1237 |
||
9055 | 1238 |
Goal "(x + x + y = y) = (x = (0::hypreal))"; |
7218 | 1239 |
by Auto_tac; |
1240 |
by (dtac (hypreal_eq_minus_iff RS iffD1) 1 THEN dtac sym 1); |
|
1241 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc])); |
|
1242 |
qed "hypreal_add_self_zero_cancel2"; |
|
1243 |
Addsimps [hypreal_add_self_zero_cancel2]; |
|
1244 |
||
9055 | 1245 |
Goal "(x + (x + y) = y) = (x = (0::hypreal))"; |
7218 | 1246 |
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1); |
1247 |
qed "hypreal_add_self_zero_cancel2a"; |
|
1248 |
Addsimps [hypreal_add_self_zero_cancel2a]; |
|
1249 |
||
1250 |
Goal "(b = -a) = (-b = (a::hypreal))"; |
|
1251 |
by Auto_tac; |
|
1252 |
qed "hypreal_minus_eq_swap"; |
|
1253 |
||
1254 |
Goal "(-b = -a) = (b = (a::hypreal))"; |
|
1255 |
by (asm_full_simp_tac (simpset() addsimps |
|
1256 |
[hypreal_minus_eq_swap]) 1); |
|
1257 |
qed "hypreal_minus_eq_cancel"; |
|
1258 |
Addsimps [hypreal_minus_eq_cancel]; |
|
1259 |
||
1260 |
Goal "x < x + 1hr"; |
|
10043 | 1261 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
1262 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add, |
|
1263 |
hypreal_one_def,hypreal_less])); |
|
7218 | 1264 |
qed "hypreal_less_self_add_one"; |
1265 |
Addsimps [hypreal_less_self_add_one]; |
|
1266 |
||
1267 |
Goal "((x::hypreal) + x = y + y) = (x = y)"; |
|
10043 | 1268 |
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); |
1269 |
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); |
|
1270 |
by (auto_tac (claset(),simpset() addsimps [hypreal_add])); |
|
1271 |
by (ALLGOALS(Ultra_tac)); |
|
7218 | 1272 |
qed "hypreal_add_self_cancel"; |
1273 |
Addsimps [hypreal_add_self_cancel]; |
|
1274 |
||
1275 |
Goal "(y = x + - y + x) = (y = (x::hypreal))"; |
|
1276 |
by Auto_tac; |
|
1277 |
by (dres_inst_tac [("x1","y")] |
|
1278 |
(hypreal_add_right_cancel RS iffD2) 1); |
|
1279 |
by (auto_tac (claset(),simpset() addsimps hypreal_add_ac)); |
|
1280 |
qed "hypreal_add_self_minus_cancel"; |
|
1281 |
Addsimps [hypreal_add_self_minus_cancel]; |
|
1282 |
||
1283 |
Goal "(y = x + (- y + x)) = (y = (x::hypreal))"; |
|
1284 |
by (asm_full_simp_tac (simpset() addsimps |
|
1285 |
[hypreal_add_assoc RS sym])1); |
|
1286 |
qed "hypreal_add_self_minus_cancel2"; |
|
1287 |
Addsimps [hypreal_add_self_minus_cancel2]; |
|
1288 |
||
10043 | 1289 |
(* of course, can prove this by "transfer" as well *) |
7218 | 1290 |
Goal "z + -x = y + (y + (-x + -z)) = (y = (z::hypreal))"; |
1291 |
by Auto_tac; |
|
1292 |
by (dres_inst_tac [("x1","z")] |
|
1293 |
(hypreal_add_right_cancel RS iffD2) 1); |
|
1294 |
by (asm_full_simp_tac (simpset() addsimps |
|
10043 | 1295 |
[hypreal_minus_add_distrib RS sym] @ hypreal_add_ac |
1296 |
delsimps [hypreal_minus_add_distrib]) 1); |
|
7218 | 1297 |
by (asm_full_simp_tac (simpset() addsimps |
1298 |
[hypreal_add_assoc RS sym,hypreal_add_right_cancel]) 1); |
|
1299 |
qed "hypreal_add_self_minus_cancel3"; |
|
1300 |
Addsimps [hypreal_add_self_minus_cancel3]; |
|
1301 |
||
9055 | 1302 |
Goal "(x * x = 0) = (x = (0::hypreal))"; |
7218 | 1303 |
by Auto_tac; |
1304 |
by (blast_tac (claset() addIs [hypreal_mult_not_0E]) 1); |
|
1305 |
qed "hypreal_mult_self_eq_zero_iff"; |
|
1306 |
Addsimps [hypreal_mult_self_eq_zero_iff]; |
|
1307 |
||
9055 | 1308 |
Goal "(0 = x * x) = (x = (0::hypreal))"; |
7218 | 1309 |
by (auto_tac (claset() addDs [sym],simpset())); |
1310 |
qed "hypreal_mult_self_eq_zero_iff2"; |
|
1311 |
Addsimps [hypreal_mult_self_eq_zero_iff2]; |
|
1312 |
||
10043 | 1313 |
Goalw [hypreal_diff_def] "(x<y) = (x-y < (0::hypreal))"; |
1314 |
by (rtac hypreal_less_minus_iff2 1); |
|
1315 |
qed "hypreal_less_eq_diff"; |
|
7218 | 1316 |
|
10043 | 1317 |
(*** Subtraction laws ***) |
7218 | 1318 |
|
10043 | 1319 |
Goal "x + (y - z) = (x + y) - (z::hypreal)"; |
1320 |
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1); |
|
1321 |
qed "hypreal_add_diff_eq"; |
|
7218 | 1322 |
|
10043 | 1323 |
Goal "(x - y) + z = (x + z) - (y::hypreal)"; |
1324 |
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1); |
|
1325 |
qed "hypreal_diff_add_eq"; |
|
7218 | 1326 |
|
10043 | 1327 |
Goal "(x - y) - z = x - (y + (z::hypreal))"; |
1328 |
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1); |
|
1329 |
qed "hypreal_diff_diff_eq"; |
|
7218 | 1330 |
|
10043 | 1331 |
Goal "x - (y - z) = (x + z) - (y::hypreal)"; |
1332 |
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1); |
|
1333 |
qed "hypreal_diff_diff_eq2"; |
|
7218 | 1334 |
|
10043 | 1335 |
Goal "(x-y < z) = (x < z + (y::hypreal))"; |
1336 |
by (stac hypreal_less_eq_diff 1); |
|
1337 |
by (res_inst_tac [("y1", "z")] (hypreal_less_eq_diff RS ssubst) 1); |
|
1338 |
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1); |
|
1339 |
qed "hypreal_diff_less_eq"; |
|
7218 | 1340 |
|
10043 | 1341 |
Goal "(x < z-y) = (x + (y::hypreal) < z)"; |
1342 |
by (stac hypreal_less_eq_diff 1); |
|
1343 |
by (res_inst_tac [("y1", "z-y")] (hypreal_less_eq_diff RS ssubst) 1); |
|
1344 |
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1); |
|
1345 |
qed "hypreal_less_diff_eq"; |
|
7218 | 1346 |
|
10043 | 1347 |
Goalw [hypreal_le_def] "(x-y <= z) = (x <= z + (y::hypreal))"; |
1348 |
by (simp_tac (simpset() addsimps [hypreal_less_diff_eq]) 1); |
|
1349 |
qed "hypreal_diff_le_eq"; |
|
7218 | 1350 |
|
10043 | 1351 |
Goalw [hypreal_le_def] "(x <= z-y) = (x + (y::hypreal) <= z)"; |
1352 |
by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1); |
|
1353 |
qed "hypreal_le_diff_eq"; |
|
7218 | 1354 |
|
10043 | 1355 |
Goalw [hypreal_diff_def] "(x-y = z) = (x = z + (y::hypreal))"; |
1356 |
by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc])); |
|
1357 |
qed "hypreal_diff_eq_eq"; |
|
7218 | 1358 |
|
10043 | 1359 |
Goalw [hypreal_diff_def] "(x = z-y) = (x + (y::hypreal) = z)"; |
1360 |
by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc])); |
|
1361 |
qed "hypreal_eq_diff_eq"; |
|
7218 | 1362 |
|
10043 | 1363 |
(*This list of rewrites simplifies (in)equalities by bringing subtractions |
1364 |
to the top and then moving negative terms to the other side. |
|
1365 |
Use with hypreal_add_ac*) |
|
1366 |
val hypreal_compare_rls = |
|
1367 |
[symmetric hypreal_diff_def, |
|
1368 |
hypreal_add_diff_eq, hypreal_diff_add_eq, hypreal_diff_diff_eq, hypreal_diff_diff_eq2, |
|
1369 |
hypreal_diff_less_eq, hypreal_less_diff_eq, hypreal_diff_le_eq, hypreal_le_diff_eq, |
|
1370 |
hypreal_diff_eq_eq, hypreal_eq_diff_eq]; |
|
7218 | 1371 |
|
1372 |
||
10043 | 1373 |
(** For the cancellation simproc. |
1374 |
The idea is to cancel like terms on opposite sides by subtraction **) |
|
7218 | 1375 |
|
10043 | 1376 |
Goal "(x::hypreal) - y = x' - y' ==> (x<y) = (x'<y')"; |
1377 |
by (stac hypreal_less_eq_diff 1); |
|
1378 |
by (res_inst_tac [("y1", "y")] (hypreal_less_eq_diff RS ssubst) 1); |
|
1379 |
by (Asm_simp_tac 1); |
|
1380 |
qed "hypreal_less_eqI"; |
|
7218 | 1381 |
|
10043 | 1382 |
Goal "(x::hypreal) - y = x' - y' ==> (y<=x) = (y'<=x')"; |
1383 |
by (dtac hypreal_less_eqI 1); |
|
1384 |
by (asm_simp_tac (simpset() addsimps [hypreal_le_def]) 1); |
|
1385 |
qed "hypreal_le_eqI"; |
|
7218 | 1386 |
|
10043 | 1387 |
Goal "(x::hypreal) - y = x' - y' ==> (x=y) = (x'=y')"; |
1388 |
by Safe_tac; |
|
1389 |
by (ALLGOALS |
|
1390 |
(asm_full_simp_tac |
|
1391 |
(simpset() addsimps [hypreal_eq_diff_eq, hypreal_diff_eq_eq]))); |
|
1392 |
qed "hypreal_eq_eqI"; |
|
7218 | 1393 |