author | wenzelm |
Wed, 03 Feb 1999 16:42:40 +0100 | |
changeset 6188 | c40e5ac04e3e |
parent 6162 | 484adda70b65 |
child 7077 | 60b098bb8b8a |
permissions | -rw-r--r-- |
5588 | 1 |
(* Title: HOL/Real/Real.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Type "real" is a linear order |
|
5078 | 7 |
*) |
8 |
||
9 |
||
10 |
||
11 |
(** lemma **) |
|
12 |
Goal "(0r < x) = (? y. x = %#y)"; |
|
13 |
by (auto_tac (claset(),simpset() addsimps [real_preal_zero_less])); |
|
14 |
by (cut_inst_tac [("x","x")] real_preal_trichotomy 1); |
|
15 |
by (blast_tac (claset() addSEs [real_less_irrefl, |
|
16 |
real_preal_not_minus_gt_zero RS notE]) 1); |
|
17 |
qed "real_gt_zero_preal_Ex"; |
|
18 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
19 |
Goal "%#z < x ==> ? y. x = %#y"; |
5078 | 20 |
by (blast_tac (claset() addSDs [real_preal_zero_less RS real_less_trans] |
21 |
addIs [real_gt_zero_preal_Ex RS iffD1]) 1); |
|
22 |
qed "real_gt_preal_preal_Ex"; |
|
23 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
24 |
Goal "%#z <= x ==> ? y. x = %#y"; |
5078 | 25 |
by (blast_tac (claset() addDs [real_le_imp_less_or_eq, |
26 |
real_gt_preal_preal_Ex]) 1); |
|
27 |
qed "real_ge_preal_preal_Ex"; |
|
28 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
29 |
Goal "y <= 0r ==> !x. y < %#x"; |
5078 | 30 |
by (auto_tac (claset() addEs [real_le_imp_less_or_eq RS disjE] |
31 |
addIs [real_preal_zero_less RSN(2,real_less_trans)], |
|
32 |
simpset() addsimps [real_preal_zero_less])); |
|
33 |
qed "real_less_all_preal"; |
|
34 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
35 |
Goal "~ 0r < y ==> !x. y < %#x"; |
5078 | 36 |
by (blast_tac (claset() addSIs [real_less_all_preal,real_leI]) 1); |
37 |
qed "real_less_all_real2"; |
|
38 |
||
5588 | 39 |
Goal "((x::real) < y) = (-y < -x)"; |
5078 | 40 |
by (rtac (real_less_sum_gt_0_iff RS subst) 1); |
41 |
by (res_inst_tac [("W1","x")] (real_less_sum_gt_0_iff RS subst) 1); |
|
42 |
by (simp_tac (simpset() addsimps [real_add_commute]) 1); |
|
43 |
qed "real_less_swap_iff"; |
|
44 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
45 |
Goal "[| R + L = S; 0r < L |] ==> R < S"; |
5078 | 46 |
by (rtac (real_less_sum_gt_0_iff RS iffD1) 1); |
5588 | 47 |
by (auto_tac (claset(), simpset() addsimps real_add_ac)); |
5078 | 48 |
qed "real_lemma_add_positive_imp_less"; |
49 |
||
50 |
Goal "!!(R::real). ? T. 0r < T & R + T = S ==> R < S"; |
|
51 |
by (blast_tac (claset() addIs [real_lemma_add_positive_imp_less]) 1); |
|
52 |
qed "real_ex_add_positive_left_less"; |
|
53 |
||
5588 | 54 |
(*Alternative definition for real_less. NOT for rewriting*) |
55 |
Goal "!!(R::real). (R < S) = (? T. 0r < T & R + T = S)"; |
|
5459 | 56 |
by (blast_tac (claset() addSIs [real_less_add_positive_left_Ex, |
5588 | 57 |
real_ex_add_positive_left_less]) 1); |
58 |
qed "real_less_iff_add"; |
|
5078 | 59 |
|
5588 | 60 |
Goal "(0r < x) = (-x < x)"; |
5459 | 61 |
by Safe_tac; |
5078 | 62 |
by (rtac ccontr 2 THEN forward_tac |
63 |
[real_leI RS real_le_imp_less_or_eq] 2); |
|
64 |
by (Step_tac 2); |
|
65 |
by (dtac (real_minus_zero_less_iff RS iffD2) 2); |
|
5459 | 66 |
by (blast_tac (claset() addIs [real_less_trans]) 2); |
5588 | 67 |
by (auto_tac (claset(), |
68 |
simpset() addsimps |
|
69 |
[real_gt_zero_preal_Ex,real_preal_minus_less_self])); |
|
5078 | 70 |
qed "real_gt_zero_iff"; |
71 |
||
5588 | 72 |
Goal "(x < 0r) = (x < -x)"; |
5078 | 73 |
by (rtac (real_minus_zero_less_iff RS subst) 1); |
74 |
by (stac real_gt_zero_iff 1); |
|
75 |
by (Full_simp_tac 1); |
|
76 |
qed "real_lt_zero_iff"; |
|
77 |
||
5588 | 78 |
Goalw [real_le_def] "(0r <= x) = (-x <= x)"; |
5078 | 79 |
by (auto_tac (claset(),simpset() addsimps [real_lt_zero_iff RS sym])); |
80 |
qed "real_ge_zero_iff"; |
|
81 |
||
5588 | 82 |
Goalw [real_le_def] "(x <= 0r) = (x <= -x)"; |
5078 | 83 |
by (auto_tac (claset(),simpset() addsimps [real_gt_zero_iff RS sym])); |
84 |
qed "real_le_zero_iff"; |
|
85 |
||
86 |
Goal "(%#m1 <= %#m2) = (m1 <= m2)"; |
|
87 |
by (auto_tac (claset() addSIs [preal_leI], |
|
88 |
simpset() addsimps [real_less_le_iff RS sym])); |
|
89 |
by (dtac preal_le_less_trans 1 THEN assume_tac 1); |
|
90 |
by (etac preal_less_irrefl 1); |
|
91 |
qed "real_preal_le_iff"; |
|
92 |
||
93 |
Goal "!!(x::real). [| 0r < x; 0r < y |] ==> 0r < x * y"; |
|
94 |
by (auto_tac (claset(),simpset() addsimps [real_gt_zero_preal_Ex])); |
|
95 |
by (res_inst_tac [("x","y*ya")] exI 1); |
|
96 |
by (full_simp_tac (simpset() addsimps [real_preal_mult]) 1); |
|
97 |
qed "real_mult_order"; |
|
98 |
||
99 |
Goal "!!(x::real). [| x < 0r; y < 0r |] ==> 0r < x * y"; |
|
100 |
by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1)); |
|
101 |
by (dtac real_mult_order 1 THEN assume_tac 1); |
|
102 |
by (Asm_full_simp_tac 1); |
|
103 |
qed "real_mult_less_zero1"; |
|
104 |
||
105 |
Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x * y"; |
|
106 |
by (REPEAT(dtac real_le_imp_less_or_eq 1)); |
|
5588 | 107 |
by (auto_tac (claset() addIs [real_mult_order, real_less_imp_le], |
108 |
simpset())); |
|
5078 | 109 |
qed "real_le_mult_order"; |
110 |
||
111 |
Goal "!!(x::real). [| x <= 0r; y <= 0r |] ==> 0r <= x * y"; |
|
112 |
by (rtac real_less_or_eq_imp_le 1); |
|
113 |
by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1); |
|
114 |
by Auto_tac; |
|
115 |
by (dtac real_le_imp_less_or_eq 1); |
|
116 |
by (auto_tac (claset() addDs [real_mult_less_zero1],simpset())); |
|
117 |
qed "real_mult_le_zero1"; |
|
118 |
||
119 |
Goal "!!(x::real). [| 0r <= x; y < 0r |] ==> x * y <= 0r"; |
|
120 |
by (rtac real_less_or_eq_imp_le 1); |
|
121 |
by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1); |
|
122 |
by Auto_tac; |
|
123 |
by (dtac (real_minus_zero_less_iff RS iffD2) 1); |
|
124 |
by (rtac (real_minus_zero_less_iff RS subst) 1); |
|
125 |
by (blast_tac (claset() addDs [real_mult_order] |
|
126 |
addIs [real_minus_mult_eq2 RS ssubst]) 1); |
|
127 |
qed "real_mult_le_zero"; |
|
128 |
||
129 |
Goal "!!(x::real). [| 0r < x; y < 0r |] ==> x*y < 0r"; |
|
130 |
by (dtac (real_minus_zero_less_iff RS iffD2) 1); |
|
131 |
by (dtac real_mult_order 1 THEN assume_tac 1); |
|
132 |
by (rtac (real_minus_zero_less_iff RS iffD1) 1); |
|
133 |
by (asm_full_simp_tac (simpset() addsimps [real_minus_mult_eq2]) 1); |
|
134 |
qed "real_mult_less_zero"; |
|
135 |
||
136 |
Goalw [real_one_def] "0r < 1r"; |
|
137 |
by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2], |
|
138 |
simpset() addsimps [real_preal_def])); |
|
139 |
qed "real_zero_less_one"; |
|
140 |
||
141 |
(*** Completeness of reals ***) |
|
142 |
(** use supremum property of preal and theorems about real_preal **) |
|
143 |
(*** a few lemmas ***) |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
144 |
Goal "! x:P. 0r < x ==> ((? x:P. y < x) = (? X. %#X : P & y < %#X))"; |
5078 | 145 |
by (blast_tac (claset() addSDs [bspec,real_gt_zero_preal_Ex RS iffD1]) 1); |
146 |
qed "real_sup_lemma1"; |
|
147 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
148 |
Goal "[| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \ |
5078 | 149 |
\ ==> (? X. X: {w. %#w : P}) & (? Y. !X: {w. %#w : P}. X < Y)"; |
150 |
by (rtac conjI 1); |
|
151 |
by (blast_tac (claset() addDs [bspec,real_gt_zero_preal_Ex RS iffD1]) 1); |
|
152 |
by Auto_tac; |
|
153 |
by (dtac bspec 1 THEN assume_tac 1); |
|
154 |
by (forward_tac [bspec] 1 THEN assume_tac 1); |
|
155 |
by (dtac real_less_trans 1 THEN assume_tac 1); |
|
156 |
by (dtac (real_gt_zero_preal_Ex RS iffD1) 1 THEN etac exE 1); |
|
157 |
by (res_inst_tac [("x","ya")] exI 1); |
|
158 |
by Auto_tac; |
|
159 |
by (dres_inst_tac [("x","%#X")] bspec 1 THEN assume_tac 1); |
|
160 |
by (etac real_preal_lessD 1); |
|
161 |
qed "real_sup_lemma2"; |
|
162 |
||
163 |
(*------------------------------------------------------------- |
|
164 |
Completeness of Positive Reals |
|
165 |
-------------------------------------------------------------*) |
|
166 |
||
167 |
(* Supremum property for the set of positive reals *) |
|
168 |
(* FIXME: long proof - can be improved - need only have one case split *) |
|
169 |
(* will do for now *) |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
170 |
Goal "[| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \ |
5078 | 171 |
\ ==> (? S. !y. (? x: P. y < x) = (y < S))"; |
172 |
by (res_inst_tac [("x","%#psup({w. %#w : P})")] exI 1); |
|
173 |
by Auto_tac; |
|
174 |
by (forward_tac [real_sup_lemma2] 1 THEN Auto_tac); |
|
175 |
by (case_tac "0r < ya" 1); |
|
176 |
by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); |
|
177 |
by (dtac real_less_all_real2 2); |
|
178 |
by Auto_tac; |
|
179 |
by (rtac (preal_complete RS spec RS iffD1) 1); |
|
180 |
by Auto_tac; |
|
181 |
by (forward_tac [real_gt_preal_preal_Ex] 1); |
|
182 |
by Auto_tac; |
|
183 |
(* second part *) |
|
184 |
by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1); |
|
185 |
by (case_tac "0r < ya" 1); |
|
186 |
by (auto_tac (claset() addSDs [real_less_all_real2, |
|
187 |
real_gt_zero_preal_Ex RS iffD1],simpset())); |
|
188 |
by (forward_tac [real_sup_lemma2] 2 THEN Auto_tac); |
|
189 |
by (forward_tac [real_sup_lemma2] 1 THEN Auto_tac); |
|
190 |
by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1); |
|
5459 | 191 |
by (Blast_tac 3); |
192 |
by (Blast_tac 1); |
|
193 |
by (Blast_tac 1); |
|
5078 | 194 |
by (Blast_tac 1); |
195 |
qed "posreal_complete"; |
|
196 |
||
5588 | 197 |
|
198 |
||
199 |
(*** Monotonicity results ***) |
|
200 |
||
201 |
Goal "(v+z < w+z) = (v < (w::real))"; |
|
202 |
by (Simp_tac 1); |
|
203 |
qed "real_add_right_cancel_less"; |
|
5078 | 204 |
|
5588 | 205 |
Goal "(z+v < z+w) = (v < (w::real))"; |
206 |
by (Simp_tac 1); |
|
207 |
qed "real_add_left_cancel_less"; |
|
208 |
||
209 |
Addsimps [real_add_right_cancel_less, real_add_left_cancel_less]; |
|
210 |
||
211 |
Goal "(v+z <= w+z) = (v <= (w::real))"; |
|
212 |
by (Simp_tac 1); |
|
213 |
qed "real_add_right_cancel_le"; |
|
5078 | 214 |
|
5588 | 215 |
Goal "(z+v <= z+w) = (v <= (w::real))"; |
216 |
by (Simp_tac 1); |
|
217 |
qed "real_add_left_cancel_le"; |
|
218 |
||
219 |
Addsimps [real_add_right_cancel_le, real_add_left_cancel_le]; |
|
220 |
||
221 |
(*"v<=w ==> v+z <= w+z"*) |
|
222 |
bind_thm ("real_add_less_mono1", real_add_right_cancel_less RS iffD2); |
|
223 |
||
224 |
(*"v<=w ==> v+z <= w+z"*) |
|
225 |
bind_thm ("real_add_le_mono1", real_add_right_cancel_le RS iffD2); |
|
226 |
||
227 |
Goal "!!z z'::real. [| w'<w; z'<=z |] ==> w' + z' < w + z"; |
|
228 |
by (etac (real_add_less_mono1 RS real_less_le_trans) 1); |
|
229 |
by (Simp_tac 1); |
|
230 |
qed "real_add_less_mono"; |
|
231 |
||
5078 | 232 |
|
233 |
Goal "!!(A::real). A < B ==> C + A < C + B"; |
|
5588 | 234 |
by (Simp_tac 1); |
5078 | 235 |
qed "real_add_less_mono2"; |
236 |
||
237 |
Goal "!!(A::real). A + C < B + C ==> A < B"; |
|
5588 | 238 |
by (Full_simp_tac 1); |
5078 | 239 |
qed "real_less_add_right_cancel"; |
240 |
||
241 |
Goal "!!(A::real). C + A < C + B ==> A < B"; |
|
5588 | 242 |
by (Full_simp_tac 1); |
5078 | 243 |
qed "real_less_add_left_cancel"; |
244 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
245 |
Goal "[| 0r < x; 0r < y |] ==> 0r < x + y"; |
6162 | 246 |
by (etac real_less_trans 1); |
247 |
by (dtac real_add_less_mono2 1); |
|
5588 | 248 |
by (Full_simp_tac 1); |
5078 | 249 |
qed "real_add_order"; |
250 |
||
251 |
Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x + y"; |
|
252 |
by (REPEAT(dtac real_le_imp_less_or_eq 1)); |
|
5588 | 253 |
by (auto_tac (claset() addIs [real_add_order, real_less_imp_le], |
254 |
simpset())); |
|
5078 | 255 |
qed "real_le_add_order"; |
256 |
||
5588 | 257 |
Goal "[| R1 < S1; R2 < S2 |] ==> R1 + R2 < S1 + (S2::real)"; |
6162 | 258 |
by (dtac real_add_less_mono1 1); |
259 |
by (etac real_less_trans 1); |
|
260 |
by (etac real_add_less_mono2 1); |
|
5078 | 261 |
qed "real_add_less_mono"; |
262 |
||
263 |
Goal "!!(q1::real). q1 <= q2 ==> x + q1 <= x + q2"; |
|
5588 | 264 |
by (Simp_tac 1); |
5078 | 265 |
qed "real_add_left_le_mono1"; |
266 |
||
5588 | 267 |
Goal "[|i<=j; k<=l |] ==> i + k <= j + (l::real)"; |
6162 | 268 |
by (dtac real_add_le_mono1 1); |
269 |
by (etac real_le_trans 1); |
|
5588 | 270 |
by (Simp_tac 1); |
5078 | 271 |
qed "real_add_le_mono"; |
272 |
||
273 |
Goal "EX (x::real). x < y"; |
|
274 |
by (rtac (real_add_zero_right RS subst) 1); |
|
5588 | 275 |
by (res_inst_tac [("x","y + -1r")] exI 1); |
5078 | 276 |
by (auto_tac (claset() addSIs [real_add_less_mono2], |
5588 | 277 |
simpset() addsimps [real_minus_zero_less_iff2, real_zero_less_one])); |
5078 | 278 |
qed "real_less_Ex"; |
5588 | 279 |
|
5078 | 280 |
(*--------------------------------------------------------------------------------- |
281 |
An embedding of the naturals in the reals |
|
282 |
---------------------------------------------------------------------------------*) |
|
283 |
||
284 |
Goalw [real_nat_def] "%%#0 = 1r"; |
|
285 |
by (full_simp_tac (simpset() addsimps [pnat_one_iff RS sym,real_preal_def]) 1); |
|
286 |
by (fold_tac [real_one_def]); |
|
287 |
by (rtac refl 1); |
|
288 |
qed "real_nat_one"; |
|
289 |
||
290 |
Goalw [real_nat_def] "%%#1 = 1r + 1r"; |
|
291 |
by (full_simp_tac (simpset() addsimps [real_preal_def,real_one_def, |
|
292 |
pnat_two_eq,real_add,prat_pnat_add RS sym,preal_prat_add RS sym |
|
293 |
] @ pnat_add_ac) 1); |
|
294 |
qed "real_nat_two"; |
|
295 |
||
296 |
Goalw [real_nat_def] |
|
297 |
"%%#n1 + %%#n2 = %%#(n1 + n2) + 1r"; |
|
298 |
by (full_simp_tac (simpset() addsimps [real_nat_one RS sym, |
|
299 |
real_nat_def,real_preal_add RS sym,preal_prat_add RS sym, |
|
300 |
prat_pnat_add RS sym,pnat_nat_add]) 1); |
|
301 |
qed "real_nat_add"; |
|
302 |
||
303 |
Goal "%%#(n + 1) = %%#n + 1r"; |
|
304 |
by (res_inst_tac [("x1","1r")] (real_add_right_cancel RS iffD1) 1); |
|
305 |
by (rtac (real_nat_add RS subst) 1); |
|
306 |
by (full_simp_tac (simpset() addsimps [real_nat_two,real_add_assoc]) 1); |
|
307 |
qed "real_nat_add_one"; |
|
308 |
||
309 |
Goal "Suc n = n + 1"; |
|
310 |
by Auto_tac; |
|
311 |
qed "lemma"; |
|
312 |
||
313 |
Goal "%%#Suc n = %%#n + 1r"; |
|
314 |
by (stac lemma 1); |
|
315 |
by (rtac real_nat_add_one 1); |
|
316 |
qed "real_nat_Suc"; |
|
317 |
||
318 |
Goal "inj(real_nat)"; |
|
319 |
by (rtac injI 1); |
|
320 |
by (rewtac real_nat_def); |
|
321 |
by (dtac (inj_real_preal RS injD) 1); |
|
322 |
by (dtac (inj_preal_prat RS injD) 1); |
|
323 |
by (dtac (inj_prat_pnat RS injD) 1); |
|
324 |
by (etac (inj_pnat_nat RS injD) 1); |
|
325 |
qed "inj_real_nat"; |
|
326 |
||
327 |
Goalw [real_nat_def] "0r < %%#n"; |
|
328 |
by (rtac (real_gt_zero_preal_Ex RS iffD2) 1); |
|
329 |
by (Blast_tac 1); |
|
330 |
qed "real_nat_less_zero"; |
|
331 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
332 |
Goal "1r <= %%#n"; |
5078 | 333 |
by (simp_tac (simpset() addsimps [real_nat_one RS sym]) 1); |
5184 | 334 |
by (induct_tac "n" 1); |
5588 | 335 |
by (auto_tac (claset(), |
336 |
simpset () addsimps [real_nat_Suc,real_nat_one, |
|
337 |
real_nat_less_zero, real_less_imp_le])); |
|
5078 | 338 |
qed "real_nat_less_one"; |
339 |
||
340 |
Goal "rinv(%%#n) ~= 0r"; |
|
341 |
by (rtac ((real_nat_less_zero RS |
|
342 |
real_not_refl2 RS not_sym) RS rinv_not_zero) 1); |
|
343 |
qed "real_nat_rinv_not_zero"; |
|
344 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
345 |
Goal "rinv(%%#x) = rinv(%%#y) ==> x = y"; |
5078 | 346 |
by (rtac (inj_real_nat RS injD) 1); |
347 |
by (res_inst_tac [("n2","x")] |
|
348 |
(real_nat_rinv_not_zero RS real_mult_left_cancel RS iffD1) 1); |
|
349 |
by (full_simp_tac (simpset() addsimps [(real_nat_less_zero RS |
|
350 |
real_not_refl2 RS not_sym) RS real_mult_inv_left]) 1); |
|
351 |
by (asm_full_simp_tac (simpset() addsimps [(real_nat_less_zero RS |
|
352 |
real_not_refl2 RS not_sym)]) 1); |
|
353 |
qed "real_nat_rinv_inj"; |
|
354 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
355 |
Goal "0r < x ==> 0r < rinv x"; |
5078 | 356 |
by (EVERY1[rtac ccontr, dtac real_leI]); |
357 |
by (forward_tac [real_minus_zero_less_iff2 RS iffD2] 1); |
|
358 |
by (forward_tac [real_not_refl2 RS not_sym] 1); |
|
359 |
by (dtac (real_not_refl2 RS not_sym RS rinv_not_zero) 1); |
|
360 |
by (EVERY1[dtac real_le_imp_less_or_eq, Step_tac]); |
|
361 |
by (dtac real_mult_less_zero1 1 THEN assume_tac 1); |
|
362 |
by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym], |
|
363 |
simpset() addsimps [real_minus_mult_eq1 RS sym])); |
|
364 |
qed "real_rinv_gt_zero"; |
|
365 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
366 |
Goal "x < 0r ==> rinv x < 0r"; |
5078 | 367 |
by (forward_tac [real_not_refl2] 1); |
368 |
by (dtac (real_minus_zero_less_iff RS iffD2) 1); |
|
369 |
by (rtac (real_minus_zero_less_iff RS iffD1) 1); |
|
370 |
by (dtac (real_minus_rinv RS sym) 1); |
|
371 |
by (auto_tac (claset() addIs [real_rinv_gt_zero], |
|
372 |
simpset())); |
|
373 |
qed "real_rinv_less_zero"; |
|
374 |
||
375 |
Goal "x+x=x*(1r+1r)"; |
|
376 |
by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1); |
|
377 |
qed "real_add_self"; |
|
378 |
||
379 |
Goal "x < x + 1r"; |
|
380 |
by (rtac (real_less_sum_gt_0_iff RS iffD1) 1); |
|
381 |
by (full_simp_tac (simpset() addsimps [real_zero_less_one, |
|
5588 | 382 |
real_add_assoc, real_add_left_commute]) 1); |
5078 | 383 |
qed "real_self_less_add_one"; |
384 |
||
385 |
Goal "1r < 1r + 1r"; |
|
386 |
by (rtac real_self_less_add_one 1); |
|
387 |
qed "real_one_less_two"; |
|
388 |
||
389 |
Goal "0r < 1r + 1r"; |
|
390 |
by (rtac ([real_zero_less_one, |
|
5588 | 391 |
real_one_less_two] MRS real_less_trans) 1); |
5078 | 392 |
qed "real_zero_less_two"; |
393 |
||
394 |
Goal "1r + 1r ~= 0r"; |
|
395 |
by (rtac (real_zero_less_two RS real_not_refl2 RS not_sym) 1); |
|
396 |
qed "real_two_not_zero"; |
|
397 |
||
398 |
Addsimps [real_two_not_zero]; |
|
399 |
||
400 |
Goal "x*rinv(1r + 1r) + x*rinv(1r + 1r) = x"; |
|
401 |
by (stac real_add_self 1); |
|
402 |
by (full_simp_tac (simpset() addsimps [real_mult_assoc]) 1); |
|
403 |
qed "real_sum_of_halves"; |
|
404 |
||
405 |
Goal "!!(x::real). [| 0r<z; x<y |] ==> x*z<y*z"; |
|
406 |
by (rotate_tac 1 1); |
|
407 |
by (dtac real_less_sum_gt_zero 1); |
|
408 |
by (rtac real_sum_gt_zero_less 1); |
|
409 |
by (dtac real_mult_order 1 THEN assume_tac 1); |
|
410 |
by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2, |
|
411 |
real_minus_mult_eq2 RS sym, real_mult_commute ]) 1); |
|
412 |
qed "real_mult_less_mono1"; |
|
413 |
||
414 |
Goal "!!(y::real). [| 0r<z; x<y |] ==> z*x<z*y"; |
|
415 |
by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_less_mono1]) 1); |
|
416 |
qed "real_mult_less_mono2"; |
|
417 |
||
418 |
Goal "!!(x::real). [| 0r<z; x*z<y*z |] ==> x<y"; |
|
419 |
by (forw_inst_tac [("x","x*z")] (real_rinv_gt_zero |
|
420 |
RS real_mult_less_mono1) 1); |
|
5588 | 421 |
by (auto_tac (claset(), |
422 |
simpset() addsimps |
|
5078 | 423 |
[real_mult_assoc,real_not_refl2 RS not_sym])); |
424 |
qed "real_mult_less_cancel1"; |
|
425 |
||
426 |
Goal "!!(x::real). [| 0r<z; z*x<z*y |] ==> x<y"; |
|
427 |
by (etac real_mult_less_cancel1 1); |
|
428 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1); |
|
429 |
qed "real_mult_less_cancel2"; |
|
430 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
431 |
Goal "0r < z ==> (x*z < y*z) = (x < y)"; |
5078 | 432 |
by (blast_tac (claset() addIs [real_mult_less_mono1, |
433 |
real_mult_less_cancel1]) 1); |
|
434 |
qed "real_mult_less_iff1"; |
|
435 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
436 |
Goal "0r < z ==> (z*x < z*y) = (x < y)"; |
5078 | 437 |
by (blast_tac (claset() addIs [real_mult_less_mono2, |
438 |
real_mult_less_cancel2]) 1); |
|
439 |
qed "real_mult_less_iff2"; |
|
440 |
||
441 |
Addsimps [real_mult_less_iff1,real_mult_less_iff2]; |
|
442 |
||
443 |
Goal "!!(x::real). [| 0r<=z; x<y |] ==> x*z<=y*z"; |
|
444 |
by (EVERY1 [rtac real_less_or_eq_imp_le, dtac real_le_imp_less_or_eq]); |
|
445 |
by (auto_tac (claset() addIs [real_mult_less_mono1],simpset())); |
|
446 |
qed "real_mult_le_less_mono1"; |
|
447 |
||
448 |
Goal "!!(x::real). [| 0r<=z; x<y |] ==> z*x<=z*y"; |
|
449 |
by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_le_less_mono1]) 1); |
|
450 |
qed "real_mult_le_less_mono2"; |
|
451 |
||
452 |
Goal "!!x y (z::real). [| 0r<=z; x<=y |] ==> z*x<=z*y"; |
|
453 |
by (dres_inst_tac [("x","x")] real_le_imp_less_or_eq 1); |
|
5588 | 454 |
by (auto_tac (claset() addIs [real_mult_le_less_mono2], simpset())); |
5078 | 455 |
qed "real_mult_le_le_mono1"; |
456 |
||
457 |
Goal "!!(x::real). x < y ==> x < (x + y)*rinv(1r + 1r)"; |
|
458 |
by (dres_inst_tac [("C","x")] real_add_less_mono2 1); |
|
459 |
by (dtac (real_add_self RS subst) 1); |
|
460 |
by (dtac (real_zero_less_two RS real_rinv_gt_zero RS |
|
461 |
real_mult_less_mono1) 1); |
|
462 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1); |
|
463 |
qed "real_less_half_sum"; |
|
464 |
||
465 |
Goal "!!(x::real). x < y ==> (x + y)*rinv(1r + 1r) < y"; |
|
5588 | 466 |
by (dtac real_add_less_mono1 1); |
5078 | 467 |
by (dtac (real_add_self RS subst) 1); |
468 |
by (dtac (real_zero_less_two RS real_rinv_gt_zero RS |
|
469 |
real_mult_less_mono1) 1); |
|
470 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1); |
|
471 |
qed "real_gt_half_sum"; |
|
472 |
||
473 |
Goal "!!(x::real). x < y ==> EX r. x < r & r < y"; |
|
474 |
by (blast_tac (claset() addSIs [real_less_half_sum,real_gt_half_sum]) 1); |
|
475 |
qed "real_dense"; |
|
476 |
||
477 |
Goal "(EX n. rinv(%%#n) < r) = (EX n. 1r < r * %%#n)"; |
|
478 |
by (Step_tac 1); |
|
479 |
by (dres_inst_tac [("n1","n")] (real_nat_less_zero |
|
480 |
RS real_mult_less_mono1) 1); |
|
481 |
by (dres_inst_tac [("n2","n")] (real_nat_less_zero RS |
|
482 |
real_rinv_gt_zero RS real_mult_less_mono1) 2); |
|
5588 | 483 |
by (auto_tac (claset(), |
484 |
simpset() addsimps [(real_nat_less_zero RS |
|
5078 | 485 |
real_not_refl2 RS not_sym),real_mult_assoc])); |
486 |
qed "real_nat_rinv_Ex_iff"; |
|
487 |
||
488 |
Goalw [real_nat_def] "(%%#n < %%#m) = (n < m)"; |
|
489 |
by Auto_tac; |
|
490 |
qed "real_nat_less_iff"; |
|
491 |
||
492 |
Addsimps [real_nat_less_iff]; |
|
493 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
494 |
Goal "0r < u ==> (u < rinv (%%#n)) = (%%#n < rinv(u))"; |
5078 | 495 |
by (Step_tac 1); |
496 |
by (res_inst_tac [("n2","n")] (real_nat_less_zero RS |
|
497 |
real_rinv_gt_zero RS real_mult_less_cancel1) 1); |
|
498 |
by (res_inst_tac [("x1","u")] ( real_rinv_gt_zero |
|
499 |
RS real_mult_less_cancel1) 2); |
|
5588 | 500 |
by (auto_tac (claset(), |
501 |
simpset() addsimps [real_nat_less_zero, |
|
5078 | 502 |
real_not_refl2 RS not_sym])); |
503 |
by (res_inst_tac [("z","u")] real_mult_less_cancel2 1); |
|
504 |
by (res_inst_tac [("n1","n")] (real_nat_less_zero RS |
|
505 |
real_mult_less_cancel2) 3); |
|
5588 | 506 |
by (auto_tac (claset(), |
507 |
simpset() addsimps [real_nat_less_zero, |
|
5078 | 508 |
real_not_refl2 RS not_sym,real_mult_assoc RS sym])); |
509 |
qed "real_nat_less_rinv_iff"; |
|
510 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
511 |
Goal "0r < u ==> (u = rinv(%%#n)) = (%%#n = rinv u)"; |
5588 | 512 |
by (auto_tac (claset(), |
513 |
simpset() addsimps [real_rinv_rinv, |
|
5078 | 514 |
real_nat_less_zero,real_not_refl2 RS not_sym])); |
515 |
qed "real_nat_rinv_eq_iff"; |
|
516 |
||
517 |
(* |
|
518 |
(*------------------------------------------------------------------ |
|
519 |
lemmas about upper bounds and least upper bound |
|
520 |
------------------------------------------------------------------*) |
|
5148
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5143
diff
changeset
|
521 |
Goalw [real_ub_def] "[| real_ub u S; x : S |] ==> x <= u"; |
5078 | 522 |
by Auto_tac; |
523 |
qed "real_ubD"; |
|
524 |
||
525 |
*) |
|
5588 | 526 |
|
527 |