author | paulson |
Tue, 02 Dec 2003 11:48:15 +0100 | |
changeset 14270 | 342451d763f9 |
parent 14269 | 502a7c95de73 |
child 14329 | ff3210fe968f |
permissions | -rw-r--r-- |
5588 | 1 |
(* Title : Real/RealDef.thy |
7219 | 2 |
ID : $Id$ |
5588 | 3 |
Author : Jacques D. Fleuriot |
4 |
Copyright : 1998 University of Cambridge |
|
5 |
Description : The reals |
|
14269 | 6 |
*) |
7 |
||
8 |
theory RealDef = PReal: |
|
5588 | 9 |
|
14270 | 10 |
(*MOVE TO THEORY PREAL*) |
14269 | 11 |
instance preal :: order |
12 |
proof qed |
|
13 |
(assumption | |
|
14 |
rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+ |
|
10752
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents:
10648
diff
changeset
|
15 |
|
14270 | 16 |
instance preal :: order |
17 |
by (intro_classes, |
|
18 |
(assumption | |
|
19 |
rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+) |
|
20 |
||
21 |
lemma preal_le_linear: "x <= y | y <= (x::preal)" |
|
22 |
apply (insert preal_linear [of x y]) |
|
23 |
apply (auto simp add: order_less_le) |
|
24 |
done |
|
25 |
||
26 |
instance preal :: linorder |
|
27 |
by (intro_classes, rule preal_le_linear) |
|
28 |
||
29 |
||
5588 | 30 |
constdefs |
31 |
realrel :: "((preal * preal) * (preal * preal)) set" |
|
14269 | 32 |
"realrel == {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" |
33 |
||
34 |
typedef (REAL) real = "UNIV//realrel" |
|
35 |
by (auto simp add: quotient_def) |
|
5588 | 36 |
|
14269 | 37 |
instance real :: ord .. |
38 |
instance real :: zero .. |
|
39 |
instance real :: one .. |
|
40 |
instance real :: plus .. |
|
41 |
instance real :: times .. |
|
42 |
instance real :: minus .. |
|
43 |
instance real :: inverse .. |
|
44 |
||
45 |
consts |
|
46 |
(*Overloaded constants denoting the Nat and Real subsets of enclosing |
|
47 |
types such as hypreal and complex*) |
|
48 |
Nats :: "'a set" |
|
49 |
Reals :: "'a set" |
|
50 |
||
51 |
(*overloaded constant for injecting other types into "real"*) |
|
52 |
real :: "'a => real" |
|
5588 | 53 |
|
54 |
||
14269 | 55 |
defs (overloaded) |
5588 | 56 |
|
14269 | 57 |
real_zero_def: |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
58 |
"0 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1), |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
59 |
preal_of_prat(prat_of_pnat 1))})" |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
60 |
|
14269 | 61 |
real_one_def: |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
62 |
"1 == Abs_REAL(realrel`` |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
63 |
{(preal_of_prat(prat_of_pnat 1) + preal_of_prat(prat_of_pnat 1), |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
64 |
preal_of_prat(prat_of_pnat 1))})" |
5588 | 65 |
|
14269 | 66 |
real_minus_def: |
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
67 |
"- R == Abs_REAL(UN (x,y):Rep_REAL(R). realrel``{(y,x)})" |
10606 | 68 |
|
14269 | 69 |
real_diff_def: |
10606 | 70 |
"R - (S::real) == R + - S" |
5588 | 71 |
|
14269 | 72 |
real_inverse_def: |
11713
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents:
11701
diff
changeset
|
73 |
"inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1)" |
5588 | 74 |
|
14269 | 75 |
real_divide_def: |
10606 | 76 |
"R / (S::real) == R * inverse S" |
14269 | 77 |
|
5588 | 78 |
constdefs |
79 |
||
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
80 |
(** these don't use the overloaded "real" function: users don't see them **) |
14269 | 81 |
|
82 |
real_of_preal :: "preal => real" |
|
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
5787
diff
changeset
|
83 |
"real_of_preal m == |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
84 |
Abs_REAL(realrel``{(m + preal_of_prat(prat_of_pnat 1), |
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
85 |
preal_of_prat(prat_of_pnat 1))})" |
5588 | 86 |
|
14269 | 87 |
real_of_posnat :: "nat => real" |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
5787
diff
changeset
|
88 |
"real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))" |
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
5787
diff
changeset
|
89 |
|
5588 | 90 |
|
14269 | 91 |
defs (overloaded) |
5588 | 92 |
|
14269 | 93 |
real_of_nat_def: "real n == real_of_posnat n + (- 1)" |
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
94 |
|
14269 | 95 |
real_add_def: |
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
96 |
"P+Q == Abs_REAL(UN p1:Rep_REAL(P). UN p2:Rep_REAL(Q). |
10834 | 97 |
(%(x1,y1). (%(x2,y2). realrel``{(x1+x2, y1+y2)}) p2) p1)" |
14269 | 98 |
|
99 |
real_mult_def: |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
100 |
"P*Q == Abs_REAL(UN p1:Rep_REAL(P). UN p2:Rep_REAL(Q). |
10834 | 101 |
(%(x1,y1). (%(x2,y2). realrel``{(x1*x2+y1*y2,x1*y2+x2*y1)}) |
10752
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents:
10648
diff
changeset
|
102 |
p2) p1)" |
5588 | 103 |
|
14269 | 104 |
real_less_def: |
105 |
"P<Q == \<exists>x1 y1 x2 y2. x1 + y2 < x2 + y1 & |
|
106 |
(x1,y1):Rep_REAL(P) & (x2,y2):Rep_REAL(Q)" |
|
107 |
real_le_def: |
|
108 |
"P \<le> (Q::real) == ~(Q < P)" |
|
5588 | 109 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12018
diff
changeset
|
110 |
syntax (xsymbols) |
14269 | 111 |
Reals :: "'a set" ("\<real>") |
112 |
Nats :: "'a set" ("\<nat>") |
|
113 |
||
114 |
||
115 |
(*** Proving that realrel is an equivalence relation ***) |
|
116 |
||
14270 | 117 |
lemma preal_trans_lemma: |
118 |
"[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] |
|
14269 | 119 |
==> x1 + y3 = x3 + y1" |
120 |
apply (rule_tac C = y2 in preal_add_right_cancel) |
|
121 |
apply (rotate_tac 1, drule sym) |
|
122 |
apply (simp add: preal_add_ac) |
|
123 |
apply (rule preal_add_left_commute [THEN subst]) |
|
124 |
apply (rule_tac x1 = x1 in preal_add_assoc [THEN subst]) |
|
125 |
apply (simp add: preal_add_ac) |
|
126 |
done |
|
127 |
||
128 |
lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)" |
|
129 |
by (unfold realrel_def, blast) |
|
130 |
||
131 |
lemma realrel_refl: "(x,x): realrel" |
|
132 |
apply (case_tac "x") |
|
133 |
apply (simp add: realrel_def) |
|
134 |
done |
|
135 |
||
136 |
lemma equiv_realrel: "equiv UNIV realrel" |
|
137 |
apply (unfold equiv_def refl_def sym_def trans_def realrel_def) |
|
138 |
apply (fast elim!: sym preal_trans_lemma) |
|
139 |
done |
|
140 |
||
141 |
(* (realrel `` {x} = realrel `` {y}) = ((x,y) : realrel) *) |
|
142 |
lemmas equiv_realrel_iff = |
|
143 |
eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I] |
|
144 |
||
145 |
declare equiv_realrel_iff [simp] |
|
146 |
||
147 |
lemma realrel_in_real [simp]: "realrel``{(x,y)}: REAL" |
|
148 |
by (unfold REAL_def realrel_def quotient_def, blast) |
|
149 |
||
150 |
lemma inj_on_Abs_REAL: "inj_on Abs_REAL REAL" |
|
151 |
apply (rule inj_on_inverseI) |
|
152 |
apply (erule Abs_REAL_inverse) |
|
153 |
done |
|
154 |
||
155 |
declare inj_on_Abs_REAL [THEN inj_on_iff, simp] |
|
156 |
declare Abs_REAL_inverse [simp] |
|
157 |
||
158 |
||
159 |
lemmas eq_realrelD = equiv_realrel [THEN [2] eq_equiv_class] |
|
160 |
||
161 |
lemma inj_Rep_REAL: "inj Rep_REAL" |
|
162 |
apply (rule inj_on_inverseI) |
|
163 |
apply (rule Rep_REAL_inverse) |
|
164 |
done |
|
165 |
||
166 |
(** real_of_preal: the injection from preal to real **) |
|
167 |
lemma inj_real_of_preal: "inj(real_of_preal)" |
|
168 |
apply (rule inj_onI) |
|
169 |
apply (unfold real_of_preal_def) |
|
170 |
apply (drule inj_on_Abs_REAL [THEN inj_onD]) |
|
171 |
apply (rule realrel_in_real)+ |
|
172 |
apply (drule eq_equiv_class) |
|
173 |
apply (rule equiv_realrel, blast) |
|
174 |
apply (simp add: realrel_def) |
|
175 |
done |
|
176 |
||
177 |
lemma eq_Abs_REAL: |
|
178 |
"(!!x y. z = Abs_REAL(realrel``{(x,y)}) ==> P) ==> P" |
|
179 |
apply (rule_tac x1 = z in Rep_REAL [unfolded REAL_def, THEN quotientE]) |
|
180 |
apply (drule_tac f = Abs_REAL in arg_cong) |
|
181 |
apply (case_tac "x") |
|
182 |
apply (simp add: Rep_REAL_inverse) |
|
183 |
done |
|
184 |
||
185 |
(**** real_minus: additive inverse on real ****) |
|
186 |
||
187 |
lemma real_minus_congruent: |
|
188 |
"congruent realrel (%p. (%(x,y). realrel``{(y,x)}) p)" |
|
189 |
apply (unfold congruent_def, clarify) |
|
190 |
apply (simp add: preal_add_commute) |
|
191 |
done |
|
192 |
||
193 |
lemma real_minus: |
|
194 |
"- (Abs_REAL(realrel``{(x,y)})) = Abs_REAL(realrel `` {(y,x)})" |
|
195 |
apply (unfold real_minus_def) |
|
196 |
apply (rule_tac f = Abs_REAL in arg_cong) |
|
197 |
apply (simp add: realrel_in_real [THEN Abs_REAL_inverse] |
|
198 |
UN_equiv_class [OF equiv_realrel real_minus_congruent]) |
|
199 |
done |
|
200 |
||
201 |
lemma real_minus_minus: "- (- z) = (z::real)" |
|
202 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
203 |
apply (simp add: real_minus) |
|
204 |
done |
|
205 |
||
206 |
declare real_minus_minus [simp] |
|
207 |
||
208 |
lemma inj_real_minus: "inj(%r::real. -r)" |
|
209 |
apply (rule inj_onI) |
|
210 |
apply (drule_tac f = uminus in arg_cong) |
|
211 |
apply (simp add: real_minus_minus) |
|
212 |
done |
|
213 |
||
214 |
lemma real_minus_zero: "- 0 = (0::real)" |
|
215 |
apply (unfold real_zero_def) |
|
216 |
apply (simp add: real_minus) |
|
217 |
done |
|
218 |
||
219 |
declare real_minus_zero [simp] |
|
220 |
||
221 |
lemma real_minus_zero_iff: "(-x = 0) = (x = (0::real))" |
|
222 |
apply (rule_tac z = x in eq_Abs_REAL) |
|
223 |
apply (auto simp add: real_zero_def real_minus preal_add_ac) |
|
224 |
done |
|
225 |
||
226 |
declare real_minus_zero_iff [simp] |
|
227 |
||
228 |
(*** Congruence property for addition ***) |
|
229 |
||
230 |
lemma real_add_congruent2_lemma: |
|
231 |
"[|a + ba = aa + b; ab + bc = ac + bb|] |
|
232 |
==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))" |
|
233 |
apply (simp add: preal_add_assoc) |
|
234 |
apply (rule preal_add_left_commute [of ab, THEN ssubst]) |
|
235 |
apply (simp add: preal_add_assoc [symmetric]) |
|
236 |
apply (simp add: preal_add_ac) |
|
237 |
done |
|
238 |
||
239 |
lemma real_add: |
|
240 |
"Abs_REAL(realrel``{(x1,y1)}) + Abs_REAL(realrel``{(x2,y2)}) = |
|
241 |
Abs_REAL(realrel``{(x1+x2, y1+y2)})" |
|
242 |
apply (simp add: real_add_def UN_UN_split_split_eq) |
|
243 |
apply (subst equiv_realrel [THEN UN_equiv_class2]) |
|
244 |
apply (auto simp add: congruent2_def) |
|
245 |
apply (blast intro: real_add_congruent2_lemma) |
|
246 |
done |
|
247 |
||
248 |
lemma real_add_commute: "(z::real) + w = w + z" |
|
249 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
250 |
apply (rule_tac z = w in eq_Abs_REAL) |
|
251 |
apply (simp add: preal_add_ac real_add) |
|
252 |
done |
|
253 |
||
254 |
lemma real_add_assoc: "((z1::real) + z2) + z3 = z1 + (z2 + z3)" |
|
255 |
apply (rule_tac z = z1 in eq_Abs_REAL) |
|
256 |
apply (rule_tac z = z2 in eq_Abs_REAL) |
|
257 |
apply (rule_tac z = z3 in eq_Abs_REAL) |
|
258 |
apply (simp add: real_add preal_add_assoc) |
|
259 |
done |
|
260 |
||
261 |
(*For AC rewriting*) |
|
262 |
lemma real_add_left_commute: "(x::real)+(y+z)=y+(x+z)" |
|
263 |
apply (rule mk_left_commute [of "op +"]) |
|
264 |
apply (rule real_add_assoc) |
|
265 |
apply (rule real_add_commute) |
|
266 |
done |
|
267 |
||
268 |
||
269 |
(* real addition is an AC operator *) |
|
270 |
lemmas real_add_ac = real_add_assoc real_add_commute real_add_left_commute |
|
271 |
||
272 |
lemma real_add_zero_left: "(0::real) + z = z" |
|
273 |
apply (unfold real_of_preal_def real_zero_def) |
|
274 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
275 |
apply (simp add: real_add preal_add_ac) |
|
276 |
done |
|
277 |
declare real_add_zero_left [simp] |
|
278 |
||
279 |
lemma real_add_zero_right: "z + (0::real) = z" |
|
280 |
by (simp add: real_add_commute) |
|
281 |
declare real_add_zero_right [simp] |
|
282 |
||
283 |
instance real :: plus_ac0 |
|
284 |
by (intro_classes, |
|
285 |
(assumption | |
|
286 |
rule real_add_commute real_add_assoc real_add_zero_left)+) |
|
287 |
||
288 |
||
289 |
lemma real_add_minus: "z + (-z) = (0::real)" |
|
290 |
apply (unfold real_zero_def) |
|
291 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
292 |
apply (simp add: real_minus real_add preal_add_commute) |
|
293 |
done |
|
294 |
declare real_add_minus [simp] |
|
295 |
||
296 |
lemma real_add_minus_left: "(-z) + z = (0::real)" |
|
297 |
by (simp add: real_add_commute) |
|
298 |
declare real_add_minus_left [simp] |
|
299 |
||
300 |
||
301 |
(*** Congruence property for multiplication ***) |
|
302 |
||
303 |
lemma real_mult_congruent2_lemma: "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> |
|
304 |
x * x1 + y * y1 + (x * y2 + x2 * y) = |
|
305 |
x * x2 + y * y2 + (x * y1 + x1 * y)" |
|
306 |
apply (simp add: preal_add_left_commute preal_add_assoc [symmetric] preal_add_mult_distrib2 [symmetric]) |
|
307 |
apply (rule preal_mult_commute [THEN subst]) |
|
308 |
apply (rule_tac y1 = x2 in preal_mult_commute [THEN subst]) |
|
309 |
apply (simp add: preal_add_assoc preal_add_mult_distrib2 [symmetric]) |
|
310 |
apply (simp add: preal_add_commute) |
|
311 |
done |
|
312 |
||
313 |
lemma real_mult_congruent2: |
|
314 |
"congruent2 realrel (%p1 p2. |
|
315 |
(%(x1,y1). (%(x2,y2). realrel``{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)" |
|
316 |
apply (rule equiv_realrel [THEN congruent2_commuteI], clarify) |
|
317 |
apply (unfold split_def) |
|
318 |
apply (simp add: preal_mult_commute preal_add_commute) |
|
319 |
apply (auto simp add: real_mult_congruent2_lemma) |
|
320 |
done |
|
321 |
||
322 |
lemma real_mult: |
|
323 |
"Abs_REAL((realrel``{(x1,y1)})) * Abs_REAL((realrel``{(x2,y2)})) = |
|
324 |
Abs_REAL(realrel `` {(x1*x2+y1*y2,x1*y2+x2*y1)})" |
|
325 |
apply (unfold real_mult_def) |
|
326 |
apply (simp add: equiv_realrel [THEN UN_equiv_class2] real_mult_congruent2) |
|
327 |
done |
|
328 |
||
329 |
lemma real_mult_commute: "(z::real) * w = w * z" |
|
330 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
331 |
apply (rule_tac z = w in eq_Abs_REAL) |
|
332 |
apply (simp add: real_mult preal_add_ac preal_mult_ac) |
|
333 |
done |
|
334 |
||
335 |
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)" |
|
336 |
apply (rule_tac z = z1 in eq_Abs_REAL) |
|
337 |
apply (rule_tac z = z2 in eq_Abs_REAL) |
|
338 |
apply (rule_tac z = z3 in eq_Abs_REAL) |
|
339 |
apply (simp add: preal_add_mult_distrib2 real_mult preal_add_ac preal_mult_ac) |
|
340 |
done |
|
341 |
||
342 |
||
343 |
(*For AC rewriting*) |
|
344 |
lemma real_mult_left_commute: "(x::real)*(y*z)=y*(x*z)" |
|
345 |
apply (rule mk_left_commute [of "op *"]) |
|
346 |
apply (rule real_mult_assoc) |
|
347 |
apply (rule real_mult_commute) |
|
348 |
done |
|
349 |
||
350 |
(* real multiplication is an AC operator *) |
|
351 |
lemmas real_mult_ac = real_mult_assoc real_mult_commute real_mult_left_commute |
|
352 |
||
353 |
lemma real_mult_1: "(1::real) * z = z" |
|
354 |
apply (unfold real_one_def pnat_one_def) |
|
355 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
356 |
apply (simp add: real_mult preal_add_mult_distrib2 preal_mult_1_right preal_mult_ac preal_add_ac) |
|
357 |
done |
|
358 |
||
359 |
declare real_mult_1 [simp] |
|
360 |
||
361 |
lemma real_mult_1_right: "z * (1::real) = z" |
|
362 |
by (simp add: real_mult_commute) |
|
363 |
||
364 |
declare real_mult_1_right [simp] |
|
365 |
||
366 |
lemma real_mult_0: "0 * z = (0::real)" |
|
367 |
apply (unfold real_zero_def pnat_one_def) |
|
368 |
apply (rule_tac z = z in eq_Abs_REAL) |
|
369 |
apply (simp add: real_mult preal_add_mult_distrib2 preal_mult_1_right preal_mult_ac preal_add_ac) |
|
370 |
done |
|
371 |
||
372 |
lemma real_mult_0_right: "z * 0 = (0::real)" |
|
373 |
by (simp add: real_mult_commute real_mult_0) |
|
374 |
||
375 |
declare real_mult_0_right [simp] real_mult_0 [simp] |
|
376 |
||
377 |
lemma real_mult_minus_eq1: "(-x) * (y::real) = -(x * y)" |
|
378 |
apply (rule_tac z = x in eq_Abs_REAL) |
|
379 |
apply (rule_tac z = y in eq_Abs_REAL) |
|
380 |
apply (auto simp add: real_minus real_mult preal_mult_ac preal_add_ac) |
|
381 |
done |
|
382 |
declare real_mult_minus_eq1 [simp] |
|
383 |
||
384 |
lemmas real_minus_mult_eq1 = real_mult_minus_eq1 [symmetric, standard] |
|
385 |
||
386 |
lemma real_mult_minus_eq2: "x * (- y :: real) = -(x * y)" |
|
387 |
by (simp add: real_mult_commute [of x]) |
|
388 |
declare real_mult_minus_eq2 [simp] |
|
389 |
||
390 |
lemmas real_minus_mult_eq2 = real_mult_minus_eq2 [symmetric, standard] |
|
391 |
||
392 |
lemma real_mult_minus_1: "(- (1::real)) * z = -z" |
|
393 |
by simp |
|
394 |
declare real_mult_minus_1 [simp] |
|
395 |
||
396 |
lemma real_mult_minus_1_right: "z * (- (1::real)) = -z" |
|
397 |
by (subst real_mult_commute, simp) |
|
398 |
declare real_mult_minus_1_right [simp] |
|
399 |
||
400 |
lemma real_minus_mult_cancel: "(-x) * (-y) = x * (y::real)" |
|
401 |
by simp |
|
402 |
||
403 |
declare real_minus_mult_cancel [simp] |
|
404 |
||
405 |
lemma real_minus_mult_commute: "(-x) * y = x * (- y :: real)" |
|
406 |
by simp |
|
407 |
||
408 |
(** Lemmas **) |
|
409 |
||
410 |
lemma real_add_assoc_cong: "(z::real) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" |
|
411 |
by (simp add: real_add_assoc [symmetric]) |
|
412 |
||
413 |
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)" |
|
414 |
apply (rule_tac z = z1 in eq_Abs_REAL) |
|
415 |
apply (rule_tac z = z2 in eq_Abs_REAL) |
|
416 |
apply (rule_tac z = w in eq_Abs_REAL) |
|
417 |
apply (simp add: preal_add_mult_distrib2 real_add real_mult preal_add_ac preal_mult_ac) |
|
418 |
done |
|
419 |
||
420 |
lemma real_add_mult_distrib2: "(w::real) * (z1 + z2) = (w * z1) + (w * z2)" |
|
421 |
by (simp add: real_mult_commute [of w] real_add_mult_distrib) |
|
422 |
||
423 |
lemma real_diff_mult_distrib: "((z1::real) - z2) * w = (z1 * w) - (z2 * w)" |
|
424 |
apply (unfold real_diff_def) |
|
425 |
apply (simp add: real_add_mult_distrib) |
|
426 |
done |
|
427 |
||
428 |
lemma real_diff_mult_distrib2: "(w::real) * (z1 - z2) = (w * z1) - (w * z2)" |
|
429 |
by (simp add: real_mult_commute [of w] real_diff_mult_distrib) |
|
430 |
||
431 |
(*** one and zero are distinct ***) |
|
432 |
lemma real_zero_not_eq_one: "0 ~= (1::real)" |
|
433 |
apply (unfold real_zero_def real_one_def) |
|
434 |
apply (auto simp add: preal_self_less_add_left [THEN preal_not_refl2]) |
|
435 |
done |
|
436 |
||
437 |
(*** existence of inverse ***) |
|
438 |
(** lemma -- alternative definition of 0 **) |
|
439 |
lemma real_zero_iff: "0 = Abs_REAL (realrel `` {(x, x)})" |
|
440 |
apply (unfold real_zero_def) |
|
441 |
apply (auto simp add: preal_add_commute) |
|
442 |
done |
|
443 |
||
444 |
lemma real_mult_inv_right_ex: |
|
445 |
"!!(x::real). x ~= 0 ==> \<exists>y. x*y = (1::real)" |
|
446 |
apply (unfold real_zero_def real_one_def) |
|
447 |
apply (rule_tac z = x in eq_Abs_REAL) |
|
448 |
apply (cut_tac x = xa and y = y in linorder_less_linear) |
|
449 |
apply (auto dest!: preal_less_add_left_Ex simp add: real_zero_iff [symmetric]) |
|
450 |
apply (rule_tac x = "Abs_REAL (realrel `` { (preal_of_prat (prat_of_pnat 1), pinv (D) + preal_of_prat (prat_of_pnat 1))}) " in exI) |
|
451 |
apply (rule_tac [2] x = "Abs_REAL (realrel `` { (pinv (D) + preal_of_prat (prat_of_pnat 1), preal_of_prat (prat_of_pnat 1))}) " in exI) |
|
452 |
apply (auto simp add: real_mult pnat_one_def preal_mult_1_right preal_add_mult_distrib2 preal_add_mult_distrib preal_mult_1 preal_mult_inv_right preal_add_ac preal_mult_ac) |
|
453 |
done |
|
454 |
||
455 |
lemma real_mult_inv_left_ex: "x ~= 0 ==> \<exists>y. y*x = (1::real)" |
|
456 |
apply (drule real_mult_inv_right_ex) |
|
457 |
apply (auto simp add: real_mult_commute) |
|
458 |
done |
|
459 |
||
460 |
lemma real_mult_inv_left: "x ~= 0 ==> inverse(x)*x = (1::real)" |
|
461 |
apply (unfold real_inverse_def) |
|
462 |
apply (frule real_mult_inv_left_ex, safe) |
|
463 |
apply (rule someI2, auto) |
|
464 |
done |
|
465 |
declare real_mult_inv_left [simp] |
|
466 |
||
467 |
lemma real_mult_inv_right: "x ~= 0 ==> x*inverse(x) = (1::real)" |
|
468 |
apply (subst real_mult_commute) |
|
469 |
apply (auto simp add: real_mult_inv_left) |
|
470 |
done |
|
471 |
declare real_mult_inv_right [simp] |
|
472 |
||
473 |
||
474 |
(*--------------------------------------------------------- |
|
475 |
Theorems for ordering |
|
476 |
--------------------------------------------------------*) |
|
477 |
(* prove introduction and elimination rules for real_less *) |
|
478 |
||
479 |
(* real_less is a strong order i.e. nonreflexive and transitive *) |
|
480 |
||
481 |
(*** lemmas ***) |
|
482 |
lemma preal_lemma_eq_rev_sum: "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y" |
|
483 |
by (simp add: preal_add_commute) |
|
484 |
||
485 |
lemma preal_add_left_commute_cancel: "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1" |
|
486 |
by (simp add: preal_add_ac) |
|
487 |
||
488 |
lemma preal_lemma_for_not_refl: "!!(x::preal). [| x + y2a = x2a + y; |
|
489 |
x + y2b = x2b + y |] |
|
490 |
==> x2a + y2b = x2b + y2a" |
|
491 |
apply (drule preal_lemma_eq_rev_sum, assumption) |
|
492 |
apply (erule_tac V = "x + y2b = x2b + y" in thin_rl) |
|
493 |
apply (simp add: preal_add_ac) |
|
494 |
apply (drule preal_add_left_commute_cancel) |
|
495 |
apply (simp add: preal_add_ac) |
|
496 |
done |
|
497 |
||
498 |
lemma real_less_not_refl: "~ (R::real) < R" |
|
499 |
apply (rule_tac z = R in eq_Abs_REAL) |
|
500 |
apply (auto simp add: real_less_def) |
|
501 |
apply (drule preal_lemma_for_not_refl, assumption, auto) |
|
502 |
done |
|
503 |
||
504 |
(*** y < y ==> P ***) |
|
505 |
lemmas real_less_irrefl = real_less_not_refl [THEN notE, standard] |
|
506 |
declare real_less_irrefl [elim!] |
|
507 |
||
508 |
lemma real_not_refl2: "!!(x::real). x < y ==> x ~= y" |
|
509 |
by (auto simp add: real_less_not_refl) |
|
510 |
||
511 |
(* lemma re-arranging and eliminating terms *) |
|
512 |
lemma preal_lemma_trans: "!! (a::preal). [| a + b = c + d; |
|
513 |
x2b + d + (c + y2e) < a + y2b + (x2e + b) |] |
|
514 |
==> x2b + y2e < x2e + y2b" |
|
515 |
apply (simp add: preal_add_ac) |
|
516 |
apply (rule_tac C = "c+d" in preal_add_left_less_cancel) |
|
517 |
apply (simp add: preal_add_assoc [symmetric]) |
|
518 |
done |
|
519 |
||
520 |
(** A MESS! heavy re-writing involved*) |
|
521 |
lemma real_less_trans: "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3" |
|
522 |
apply (rule_tac z = R1 in eq_Abs_REAL) |
|
523 |
apply (rule_tac z = R2 in eq_Abs_REAL) |
|
524 |
apply (rule_tac z = R3 in eq_Abs_REAL) |
|
525 |
apply (auto simp add: real_less_def) |
|
526 |
apply (rule exI)+ |
|
527 |
apply (rule conjI, rule_tac [2] conjI) |
|
528 |
prefer 2 apply blast |
|
529 |
prefer 2 apply blast |
|
530 |
apply (drule preal_lemma_for_not_refl, assumption) |
|
531 |
apply (blast dest: preal_add_less_mono intro: preal_lemma_trans) |
|
532 |
done |
|
533 |
||
534 |
lemma real_less_not_sym: "!! (R1::real). R1 < R2 ==> ~ (R2 < R1)" |
|
535 |
apply (rule notI) |
|
536 |
apply (drule real_less_trans, assumption) |
|
537 |
apply (simp add: real_less_not_refl) |
|
538 |
done |
|
539 |
||
540 |
(* [| x < y; ~P ==> y < x |] ==> P *) |
|
541 |
lemmas real_less_asym = real_less_not_sym [THEN contrapos_np, standard] |
|
542 |
||
543 |
lemma real_of_preal_add: |
|
544 |
"real_of_preal ((z1::preal) + z2) = |
|
545 |
real_of_preal z1 + real_of_preal z2" |
|
546 |
apply (unfold real_of_preal_def) |
|
547 |
apply (simp add: real_add preal_add_mult_distrib preal_mult_1 add: preal_add_ac) |
|
548 |
done |
|
549 |
||
550 |
lemma real_of_preal_mult: |
|
551 |
"real_of_preal ((z1::preal) * z2) = |
|
552 |
real_of_preal z1* real_of_preal z2" |
|
553 |
apply (unfold real_of_preal_def) |
|
554 |
apply (simp (no_asm_use) add: real_mult preal_add_mult_distrib2 preal_mult_1 preal_mult_1_right pnat_one_def preal_add_ac preal_mult_ac) |
|
555 |
done |
|
556 |
||
557 |
lemma real_of_preal_ExI: |
|
558 |
"!!(x::preal). y < x ==> |
|
559 |
\<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m" |
|
560 |
apply (unfold real_of_preal_def) |
|
561 |
apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_ac) |
|
562 |
done |
|
563 |
||
564 |
lemma real_of_preal_ExD: |
|
565 |
"!!(x::preal). \<exists>m. Abs_REAL (realrel `` {(x,y)}) = |
|
566 |
real_of_preal m ==> y < x" |
|
567 |
apply (unfold real_of_preal_def) |
|
568 |
apply (auto simp add: preal_add_commute preal_add_assoc) |
|
569 |
apply (simp add: preal_add_assoc [symmetric] preal_self_less_add_left) |
|
570 |
done |
|
571 |
||
572 |
lemma real_of_preal_iff: "(\<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m) = (y < x)" |
|
573 |
by (blast intro!: real_of_preal_ExI real_of_preal_ExD) |
|
574 |
||
575 |
(*** Gleason prop 9-4.4 p 127 ***) |
|
576 |
lemma real_of_preal_trichotomy: |
|
577 |
"\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)" |
|
578 |
apply (unfold real_of_preal_def real_zero_def) |
|
579 |
apply (rule_tac z = x in eq_Abs_REAL) |
|
580 |
apply (auto simp add: real_minus preal_add_ac) |
|
581 |
apply (cut_tac x = x and y = y in linorder_less_linear) |
|
582 |
apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_assoc [symmetric]) |
|
583 |
apply (auto simp add: preal_add_commute) |
|
584 |
done |
|
585 |
||
586 |
lemma real_of_preal_trichotomyE: "!!P. [| !!m. x = real_of_preal m ==> P; |
|
587 |
x = 0 ==> P; |
|
588 |
!!m. x = -(real_of_preal m) ==> P |] ==> P" |
|
589 |
apply (cut_tac x = x in real_of_preal_trichotomy, auto) |
|
590 |
done |
|
591 |
||
592 |
lemma real_of_preal_lessD: |
|
593 |
"real_of_preal m1 < real_of_preal m2 ==> m1 < m2" |
|
594 |
apply (unfold real_of_preal_def) |
|
595 |
apply (auto simp add: real_less_def preal_add_ac) |
|
596 |
apply (auto simp add: preal_add_assoc [symmetric]) |
|
597 |
apply (auto simp add: preal_add_ac) |
|
598 |
done |
|
599 |
||
600 |
lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2" |
|
601 |
apply (drule preal_less_add_left_Ex) |
|
602 |
apply (auto simp add: real_of_preal_add real_of_preal_def real_less_def) |
|
603 |
apply (rule exI)+ |
|
604 |
apply (rule conjI, rule_tac [2] conjI) |
|
605 |
apply (rule_tac [2] refl)+ |
|
606 |
apply (simp add: preal_self_less_add_left del: preal_add_less_iff2) |
|
607 |
done |
|
608 |
||
609 |
lemma real_of_preal_less_iff1: "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)" |
|
610 |
by (blast intro: real_of_preal_lessI real_of_preal_lessD) |
|
611 |
||
612 |
declare real_of_preal_less_iff1 [simp] |
|
613 |
||
614 |
lemma real_of_preal_minus_less_self: "- real_of_preal m < real_of_preal m" |
|
615 |
apply (auto simp add: real_of_preal_def real_less_def real_minus) |
|
616 |
apply (rule exI)+ |
|
617 |
apply (rule conjI, rule_tac [2] conjI) |
|
618 |
apply (rule_tac [2] refl)+ |
|
619 |
apply (simp (no_asm_use) add: preal_add_ac) |
|
620 |
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric]) |
|
621 |
done |
|
622 |
||
623 |
lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0" |
|
624 |
apply (unfold real_zero_def) |
|
625 |
apply (auto simp add: real_of_preal_def real_less_def real_minus) |
|
626 |
apply (rule exI)+ |
|
627 |
apply (rule conjI, rule_tac [2] conjI) |
|
628 |
apply (rule_tac [2] refl)+ |
|
629 |
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac) |
|
630 |
done |
|
631 |
||
632 |
lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m" |
|
633 |
apply (cut_tac real_of_preal_minus_less_zero) |
|
634 |
apply (fast dest: real_less_trans elim: real_less_irrefl) |
|
635 |
done |
|
636 |
||
637 |
lemma real_of_preal_zero_less: "0 < real_of_preal m" |
|
638 |
apply (unfold real_zero_def) |
|
639 |
apply (auto simp add: real_of_preal_def real_less_def real_minus) |
|
640 |
apply (rule exI)+ |
|
641 |
apply (rule conjI, rule_tac [2] conjI) |
|
642 |
apply (rule_tac [2] refl)+ |
|
643 |
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac) |
|
644 |
done |
|
645 |
||
646 |
lemma real_of_preal_not_less_zero: "~ real_of_preal m < 0" |
|
647 |
apply (cut_tac real_of_preal_zero_less) |
|
648 |
apply (blast dest: real_less_trans elim: real_less_irrefl) |
|
649 |
done |
|
650 |
||
651 |
lemma real_minus_minus_zero_less: "0 < - (- real_of_preal m)" |
|
652 |
by (simp add: real_of_preal_zero_less) |
|
653 |
||
654 |
(* another lemma *) |
|
655 |
lemma real_of_preal_sum_zero_less: |
|
656 |
"0 < real_of_preal m + real_of_preal m1" |
|
657 |
apply (unfold real_zero_def) |
|
658 |
apply (auto simp add: real_of_preal_def real_less_def real_add) |
|
659 |
apply (rule exI)+ |
|
660 |
apply (rule conjI, rule_tac [2] conjI) |
|
661 |
apply (rule_tac [2] refl)+ |
|
662 |
apply (simp (no_asm_use) add: preal_add_ac) |
|
663 |
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric]) |
|
664 |
done |
|
665 |
||
666 |
lemma real_of_preal_minus_less_all: "- real_of_preal m < real_of_preal m1" |
|
667 |
apply (auto simp add: real_of_preal_def real_less_def real_minus) |
|
668 |
apply (rule exI)+ |
|
669 |
apply (rule conjI, rule_tac [2] conjI) |
|
670 |
apply (rule_tac [2] refl)+ |
|
671 |
apply (simp (no_asm_use) add: preal_add_ac) |
|
672 |
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric]) |
|
673 |
done |
|
674 |
||
675 |
lemma real_of_preal_not_minus_gt_all: "~ real_of_preal m < - real_of_preal m1" |
|
676 |
apply (cut_tac real_of_preal_minus_less_all) |
|
677 |
apply (blast dest: real_less_trans elim: real_less_irrefl) |
|
678 |
done |
|
679 |
||
680 |
lemma real_of_preal_minus_less_rev1: "- real_of_preal m1 < - real_of_preal m2 |
|
681 |
==> real_of_preal m2 < real_of_preal m1" |
|
682 |
apply (auto simp add: real_of_preal_def real_less_def real_minus) |
|
683 |
apply (rule exI)+ |
|
684 |
apply (rule conjI, rule_tac [2] conjI) |
|
685 |
apply (rule_tac [2] refl)+ |
|
686 |
apply (auto simp add: preal_add_ac) |
|
687 |
apply (simp add: preal_add_assoc [symmetric]) |
|
688 |
apply (auto simp add: preal_add_ac) |
|
689 |
done |
|
690 |
||
691 |
lemma real_of_preal_minus_less_rev2: "real_of_preal m1 < real_of_preal m2 |
|
692 |
==> - real_of_preal m2 < - real_of_preal m1" |
|
693 |
apply (auto simp add: real_of_preal_def real_less_def real_minus) |
|
694 |
apply (rule exI)+ |
|
695 |
apply (rule conjI, rule_tac [2] conjI) |
|
696 |
apply (rule_tac [2] refl)+ |
|
697 |
apply (auto simp add: preal_add_ac) |
|
698 |
apply (simp add: preal_add_assoc [symmetric]) |
|
699 |
apply (auto simp add: preal_add_ac) |
|
700 |
done |
|
701 |
||
702 |
lemma real_of_preal_minus_less_rev_iff: "(- real_of_preal m1 < - real_of_preal m2) = |
|
703 |
(real_of_preal m2 < real_of_preal m1)" |
|
704 |
apply (blast intro!: real_of_preal_minus_less_rev1 real_of_preal_minus_less_rev2) |
|
705 |
done |
|
706 |
||
707 |
declare real_of_preal_minus_less_rev_iff [simp] |
|
708 |
||
14270 | 709 |
|
710 |
subsection{*Linearity of the Ordering*} |
|
711 |
||
14269 | 712 |
lemma real_linear: "(x::real) < y | x = y | y < x" |
713 |
apply (rule_tac x = x in real_of_preal_trichotomyE) |
|
714 |
apply (rule_tac [!] x = y in real_of_preal_trichotomyE) |
|
14270 | 715 |
apply (auto dest!: preal_le_anti_sym |
716 |
simp add: preal_less_le_iff real_of_preal_minus_less_zero |
|
717 |
real_of_preal_zero_less real_of_preal_minus_less_all) |
|
14269 | 718 |
done |
719 |
||
720 |
lemma real_neq_iff: "!!w::real. (w ~= z) = (w<z | z<w)" |
|
721 |
by (cut_tac real_linear, blast) |
|
722 |
||
723 |
||
724 |
lemma real_linear_less2: "!!(R1::real). [| R1 < R2 ==> P; R1 = R2 ==> P; |
|
725 |
R2 < R1 ==> P |] ==> P" |
|
726 |
apply (cut_tac x = R1 and y = R2 in real_linear, auto) |
|
727 |
done |
|
728 |
||
729 |
lemma real_minus_zero_less_iff: "(0 < -R) = (R < (0::real))" |
|
730 |
apply (rule_tac x = R in real_of_preal_trichotomyE) |
|
731 |
apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero) |
|
732 |
done |
|
733 |
declare real_minus_zero_less_iff [simp] |
|
734 |
||
735 |
lemma real_minus_zero_less_iff2: "(-R < 0) = ((0::real) < R)" |
|
736 |
apply (rule_tac x = R in real_of_preal_trichotomyE) |
|
737 |
apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero) |
|
738 |
done |
|
739 |
declare real_minus_zero_less_iff2 [simp] |
|
740 |
||
741 |
ML |
|
742 |
{* |
|
743 |
val real_le_def = thm "real_le_def"; |
|
744 |
val real_diff_def = thm "real_diff_def"; |
|
745 |
val real_divide_def = thm "real_divide_def"; |
|
746 |
val real_of_nat_def = thm "real_of_nat_def"; |
|
747 |
||
748 |
val preal_trans_lemma = thm"preal_trans_lemma"; |
|
749 |
val realrel_iff = thm"realrel_iff"; |
|
750 |
val realrel_refl = thm"realrel_refl"; |
|
751 |
val equiv_realrel = thm"equiv_realrel"; |
|
752 |
val equiv_realrel_iff = thm"equiv_realrel_iff"; |
|
753 |
val realrel_in_real = thm"realrel_in_real"; |
|
754 |
val inj_on_Abs_REAL = thm"inj_on_Abs_REAL"; |
|
755 |
val eq_realrelD = thm"eq_realrelD"; |
|
756 |
val inj_Rep_REAL = thm"inj_Rep_REAL"; |
|
757 |
val inj_real_of_preal = thm"inj_real_of_preal"; |
|
758 |
val eq_Abs_REAL = thm"eq_Abs_REAL"; |
|
759 |
val real_minus_congruent = thm"real_minus_congruent"; |
|
760 |
val real_minus = thm"real_minus"; |
|
761 |
val real_minus_minus = thm"real_minus_minus"; |
|
762 |
val inj_real_minus = thm"inj_real_minus"; |
|
763 |
val real_minus_zero = thm"real_minus_zero"; |
|
764 |
val real_minus_zero_iff = thm"real_minus_zero_iff"; |
|
765 |
val real_add_congruent2_lemma = thm"real_add_congruent2_lemma"; |
|
766 |
val real_add = thm"real_add"; |
|
767 |
val real_add_commute = thm"real_add_commute"; |
|
768 |
val real_add_assoc = thm"real_add_assoc"; |
|
769 |
val real_add_left_commute = thm"real_add_left_commute"; |
|
770 |
val real_add_zero_left = thm"real_add_zero_left"; |
|
771 |
val real_add_zero_right = thm"real_add_zero_right"; |
|
772 |
val real_add_minus = thm"real_add_minus"; |
|
773 |
val real_add_minus_left = thm"real_add_minus_left"; |
|
774 |
||
775 |
val real_add_ac = thms"real_add_ac"; |
|
776 |
val real_mult_ac = thms"real_mult_ac"; |
|
777 |
*} |
|
778 |
||
10752
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents:
10648
diff
changeset
|
779 |
|
5588 | 780 |
end |