author | berghofe |
Fri, 01 Jul 2005 13:54:12 +0200 | |
changeset 16633 | 208ebc9311f2 |
parent 16417 | 9bc16273c2d4 |
child 16819 | 00d8f9300d13 |
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 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
5 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 |
14269 | 6 |
*) |
7 |
||
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
8 |
header{*Defining the Reals from the Positive Reals*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
9 |
|
15131 | 10 |
theory RealDef |
15140 | 11 |
imports PReal |
16417 | 12 |
uses ("real_arith.ML") |
15131 | 13 |
begin |
5588 | 14 |
|
15 |
constdefs |
|
16 |
realrel :: "((preal * preal) * (preal * preal)) set" |
|
14269 | 17 |
"realrel == {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" |
18 |
||
14484 | 19 |
typedef (Real) real = "UNIV//realrel" |
14269 | 20 |
by (auto simp add: quotient_def) |
5588 | 21 |
|
14691 | 22 |
instance real :: "{ord, zero, one, plus, times, minus, inverse}" .. |
14269 | 23 |
|
14484 | 24 |
constdefs |
25 |
||
26 |
(** these don't use the overloaded "real" function: users don't see them **) |
|
27 |
||
28 |
real_of_preal :: "preal => real" |
|
29 |
"real_of_preal m == |
|
30 |
Abs_Real(realrel``{(m + preal_of_rat 1, preal_of_rat 1)})" |
|
31 |
||
14269 | 32 |
consts |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
33 |
(*Overloaded constant denoting the Real subset of enclosing |
14269 | 34 |
types such as hypreal and complex*) |
35 |
Reals :: "'a set" |
|
36 |
||
37 |
(*overloaded constant for injecting other types into "real"*) |
|
38 |
real :: "'a => real" |
|
5588 | 39 |
|
14691 | 40 |
syntax (xsymbols) |
41 |
Reals :: "'a set" ("\<real>") |
|
42 |
||
5588 | 43 |
|
14269 | 44 |
defs (overloaded) |
5588 | 45 |
|
14269 | 46 |
real_zero_def: |
14484 | 47 |
"0 == Abs_Real(realrel``{(preal_of_rat 1, preal_of_rat 1)})" |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
48 |
|
14269 | 49 |
real_one_def: |
14484 | 50 |
"1 == Abs_Real(realrel`` |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
51 |
{(preal_of_rat 1 + preal_of_rat 1, |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
52 |
preal_of_rat 1)})" |
5588 | 53 |
|
14269 | 54 |
real_minus_def: |
14484 | 55 |
"- r == contents (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })" |
56 |
||
57 |
real_add_def: |
|
58 |
"z + w == |
|
59 |
contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w). |
|
60 |
{ Abs_Real(realrel``{(x+u, y+v)}) })" |
|
10606 | 61 |
|
14269 | 62 |
real_diff_def: |
14484 | 63 |
"r - (s::real) == r + - s" |
64 |
||
65 |
real_mult_def: |
|
66 |
"z * w == |
|
67 |
contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w). |
|
68 |
{ Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })" |
|
5588 | 69 |
|
14269 | 70 |
real_inverse_def: |
11713
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents:
11701
diff
changeset
|
71 |
"inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1)" |
5588 | 72 |
|
14269 | 73 |
real_divide_def: |
10606 | 74 |
"R / (S::real) == R * inverse S" |
14269 | 75 |
|
14484 | 76 |
real_le_def: |
77 |
"z \<le> (w::real) == |
|
78 |
\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w" |
|
5588 | 79 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
80 |
real_less_def: "(x < (y::real)) == (x \<le> y & x \<noteq> y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
81 |
|
14334 | 82 |
real_abs_def: "abs (r::real) == (if 0 \<le> r then r else -r)" |
83 |
||
84 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
85 |
|
14329 | 86 |
subsection{*Proving that realrel is an equivalence relation*} |
14269 | 87 |
|
14270 | 88 |
lemma preal_trans_lemma: |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
89 |
assumes "x + y1 = x1 + y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
90 |
and "x + y2 = x2 + y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
91 |
shows "x1 + y2 = x2 + (y1::preal)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
92 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
93 |
have "(x1 + y2) + x = (x + y2) + x1" by (simp add: preal_add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
94 |
also have "... = (x2 + y) + x1" by (simp add: prems) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
95 |
also have "... = x2 + (x1 + y)" by (simp add: preal_add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
96 |
also have "... = x2 + (x + y1)" by (simp add: prems) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
97 |
also have "... = (x2 + y1) + x" by (simp add: preal_add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
98 |
finally have "(x1 + y2) + x = (x2 + y1) + x" . |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
99 |
thus ?thesis by (simp add: preal_add_right_cancel_iff) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
100 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
101 |
|
14269 | 102 |
|
14484 | 103 |
lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)" |
104 |
by (simp add: realrel_def) |
|
14269 | 105 |
|
106 |
lemma equiv_realrel: "equiv UNIV realrel" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
107 |
apply (auto simp add: equiv_def refl_def sym_def trans_def realrel_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
108 |
apply (blast dest: preal_trans_lemma) |
14269 | 109 |
done |
110 |
||
14497 | 111 |
text{*Reduces equality of equivalence classes to the @{term realrel} relation: |
112 |
@{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *} |
|
14269 | 113 |
lemmas equiv_realrel_iff = |
114 |
eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I] |
|
115 |
||
116 |
declare equiv_realrel_iff [simp] |
|
117 |
||
14497 | 118 |
|
14484 | 119 |
lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real" |
120 |
by (simp add: Real_def realrel_def quotient_def, blast) |
|
14269 | 121 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
122 |
|
14484 | 123 |
lemma inj_on_Abs_Real: "inj_on Abs_Real Real" |
14269 | 124 |
apply (rule inj_on_inverseI) |
14484 | 125 |
apply (erule Abs_Real_inverse) |
14269 | 126 |
done |
127 |
||
14484 | 128 |
declare inj_on_Abs_Real [THEN inj_on_iff, simp] |
129 |
declare Abs_Real_inverse [simp] |
|
14269 | 130 |
|
131 |
||
14484 | 132 |
text{*Case analysis on the representation of a real number as an equivalence |
133 |
class of pairs of positive reals.*} |
|
134 |
lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: |
|
135 |
"(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P" |
|
136 |
apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE]) |
|
137 |
apply (drule arg_cong [where f=Abs_Real]) |
|
138 |
apply (auto simp add: Rep_Real_inverse) |
|
14269 | 139 |
done |
140 |
||
141 |
||
14329 | 142 |
subsection{*Congruence property for addition*} |
14269 | 143 |
|
144 |
lemma real_add_congruent2_lemma: |
|
145 |
"[|a + ba = aa + b; ab + bc = ac + bb|] |
|
146 |
==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))" |
|
147 |
apply (simp add: preal_add_assoc) |
|
148 |
apply (rule preal_add_left_commute [of ab, THEN ssubst]) |
|
149 |
apply (simp add: preal_add_assoc [symmetric]) |
|
150 |
apply (simp add: preal_add_ac) |
|
151 |
done |
|
152 |
||
153 |
lemma real_add: |
|
14497 | 154 |
"Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) = |
155 |
Abs_Real (realrel``{(x+u, y+v)})" |
|
156 |
proof - |
|
15169 | 157 |
have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z) |
158 |
respects2 realrel" |
|
14497 | 159 |
by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma) |
160 |
thus ?thesis |
|
161 |
by (simp add: real_add_def UN_UN_split_split_eq |
|
14658 | 162 |
UN_equiv_class2 [OF equiv_realrel equiv_realrel]) |
14497 | 163 |
qed |
14269 | 164 |
|
165 |
lemma real_add_commute: "(z::real) + w = w + z" |
|
14497 | 166 |
by (cases z, cases w, simp add: real_add preal_add_ac) |
14269 | 167 |
|
168 |
lemma real_add_assoc: "((z1::real) + z2) + z3 = z1 + (z2 + z3)" |
|
14497 | 169 |
by (cases z1, cases z2, cases z3, simp add: real_add preal_add_assoc) |
14269 | 170 |
|
171 |
lemma real_add_zero_left: "(0::real) + z = z" |
|
14497 | 172 |
by (cases z, simp add: real_add real_zero_def preal_add_ac) |
14269 | 173 |
|
14738 | 174 |
instance real :: comm_monoid_add |
14269 | 175 |
by (intro_classes, |
176 |
(assumption | |
|
177 |
rule real_add_commute real_add_assoc real_add_zero_left)+) |
|
178 |
||
179 |
||
14334 | 180 |
subsection{*Additive Inverse on real*} |
181 |
||
14484 | 182 |
lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})" |
183 |
proof - |
|
15169 | 184 |
have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel" |
14484 | 185 |
by (simp add: congruent_def preal_add_commute) |
186 |
thus ?thesis |
|
187 |
by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel]) |
|
188 |
qed |
|
14334 | 189 |
|
190 |
lemma real_add_minus_left: "(-z) + z = (0::real)" |
|
14497 | 191 |
by (cases z, simp add: real_minus real_add real_zero_def preal_add_commute) |
14269 | 192 |
|
193 |
||
14329 | 194 |
subsection{*Congruence property for multiplication*} |
14269 | 195 |
|
14329 | 196 |
lemma real_mult_congruent2_lemma: |
197 |
"!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> |
|
14484 | 198 |
x * x1 + y * y1 + (x * y2 + y * x2) = |
199 |
x * x2 + y * y2 + (x * y1 + y * x1)" |
|
200 |
apply (simp add: preal_add_left_commute preal_add_assoc [symmetric]) |
|
14269 | 201 |
apply (simp add: preal_add_assoc preal_add_mult_distrib2 [symmetric]) |
202 |
apply (simp add: preal_add_commute) |
|
203 |
done |
|
204 |
||
205 |
lemma real_mult_congruent2: |
|
15169 | 206 |
"(%p1 p2. |
14484 | 207 |
(%(x1,y1). (%(x2,y2). |
15169 | 208 |
{ Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1) |
209 |
respects2 realrel" |
|
14658 | 210 |
apply (rule congruent2_commuteI [OF equiv_realrel], clarify) |
14269 | 211 |
apply (simp add: preal_mult_commute preal_add_commute) |
212 |
apply (auto simp add: real_mult_congruent2_lemma) |
|
213 |
done |
|
214 |
||
215 |
lemma real_mult: |
|
14484 | 216 |
"Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) = |
217 |
Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})" |
|
218 |
by (simp add: real_mult_def UN_UN_split_split_eq |
|
14658 | 219 |
UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2]) |
14269 | 220 |
|
221 |
lemma real_mult_commute: "(z::real) * w = w * z" |
|
14497 | 222 |
by (cases z, cases w, simp add: real_mult preal_add_ac preal_mult_ac) |
14269 | 223 |
|
224 |
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)" |
|
14484 | 225 |
apply (cases z1, cases z2, cases z3) |
226 |
apply (simp add: real_mult preal_add_mult_distrib2 preal_add_ac preal_mult_ac) |
|
14269 | 227 |
done |
228 |
||
229 |
lemma real_mult_1: "(1::real) * z = z" |
|
14484 | 230 |
apply (cases z) |
231 |
apply (simp add: real_mult real_one_def preal_add_mult_distrib2 |
|
232 |
preal_mult_1_right preal_mult_ac preal_add_ac) |
|
14269 | 233 |
done |
234 |
||
235 |
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)" |
|
14484 | 236 |
apply (cases z1, cases z2, cases w) |
237 |
apply (simp add: real_add real_mult preal_add_mult_distrib2 |
|
238 |
preal_add_ac preal_mult_ac) |
|
14269 | 239 |
done |
240 |
||
14329 | 241 |
text{*one and zero are distinct*} |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
242 |
lemma real_zero_not_eq_one: "0 \<noteq> (1::real)" |
14484 | 243 |
proof - |
244 |
have "preal_of_rat 1 < preal_of_rat 1 + preal_of_rat 1" |
|
245 |
by (simp add: preal_self_less_add_left) |
|
246 |
thus ?thesis |
|
247 |
by (simp add: real_zero_def real_one_def preal_add_right_cancel_iff) |
|
248 |
qed |
|
14269 | 249 |
|
14329 | 250 |
subsection{*existence of inverse*} |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
251 |
|
14484 | 252 |
lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0" |
14497 | 253 |
by (simp add: real_zero_def preal_add_commute) |
14269 | 254 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
255 |
text{*Instead of using an existential quantifier and constructing the inverse |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
256 |
within the proof, we could define the inverse explicitly.*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
257 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
258 |
lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)" |
14484 | 259 |
apply (simp add: real_zero_def real_one_def, cases x) |
14269 | 260 |
apply (cut_tac x = xa and y = y in linorder_less_linear) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
261 |
apply (auto dest!: less_add_left_Ex simp add: real_zero_iff) |
14334 | 262 |
apply (rule_tac |
14484 | 263 |
x = "Abs_Real (realrel `` { (preal_of_rat 1, |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
264 |
inverse (D) + preal_of_rat 1)}) " |
14334 | 265 |
in exI) |
266 |
apply (rule_tac [2] |
|
14484 | 267 |
x = "Abs_Real (realrel `` { (inverse (D) + preal_of_rat 1, |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
268 |
preal_of_rat 1)})" |
14334 | 269 |
in exI) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
270 |
apply (auto simp add: real_mult preal_mult_1_right |
14329 | 271 |
preal_add_mult_distrib2 preal_add_mult_distrib preal_mult_1 |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
272 |
preal_mult_inverse_right preal_add_ac preal_mult_ac) |
14269 | 273 |
done |
274 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
275 |
lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)" |
14484 | 276 |
apply (simp add: real_inverse_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
277 |
apply (frule real_mult_inverse_left_ex, safe) |
14269 | 278 |
apply (rule someI2, auto) |
279 |
done |
|
14334 | 280 |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
281 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
282 |
subsection{*The Real Numbers form a Field*} |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
283 |
|
14334 | 284 |
instance real :: field |
285 |
proof |
|
286 |
fix x y z :: real |
|
287 |
show "- x + x = 0" by (rule real_add_minus_left) |
|
288 |
show "x - y = x + (-y)" by (simp add: real_diff_def) |
|
289 |
show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc) |
|
290 |
show "x * y = y * x" by (rule real_mult_commute) |
|
291 |
show "1 * x = x" by (rule real_mult_1) |
|
292 |
show "(x + y) * z = x * z + y * z" by (simp add: real_add_mult_distrib) |
|
293 |
show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one) |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
294 |
show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left) |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14426
diff
changeset
|
295 |
show "x / y = x * inverse y" by (simp add: real_divide_def) |
14334 | 296 |
qed |
297 |
||
298 |
||
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
299 |
text{*Inverse of zero! Useful to simplify certain equations*} |
14269 | 300 |
|
14334 | 301 |
lemma INVERSE_ZERO: "inverse 0 = (0::real)" |
14484 | 302 |
by (simp add: real_inverse_def) |
14334 | 303 |
|
304 |
instance real :: division_by_zero |
|
305 |
proof |
|
306 |
show "inverse 0 = (0::real)" by (rule INVERSE_ZERO) |
|
307 |
qed |
|
308 |
||
309 |
||
310 |
(*Pull negations out*) |
|
311 |
declare minus_mult_right [symmetric, simp] |
|
312 |
minus_mult_left [symmetric, simp] |
|
313 |
||
314 |
lemma real_mult_1_right: "z * (1::real) = z" |
|
14738 | 315 |
by (rule OrderedGroup.mult_1_right) |
14269 | 316 |
|
317 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
318 |
subsection{*The @{text "\<le>"} Ordering*} |
14269 | 319 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
320 |
lemma real_le_refl: "w \<le> (w::real)" |
14484 | 321 |
by (cases w, force simp add: real_le_def) |
14269 | 322 |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
323 |
text{*The arithmetic decision procedure is not set up for type preal. |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
324 |
This lemma is currently unused, but it could simplify the proofs of the |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
325 |
following two lemmas.*} |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
326 |
lemma preal_eq_le_imp_le: |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
327 |
assumes eq: "a+b = c+d" and le: "c \<le> a" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
328 |
shows "b \<le> (d::preal)" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
329 |
proof - |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
330 |
have "c+d \<le> a+d" by (simp add: prems preal_cancels) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
331 |
hence "a+b \<le> a+d" by (simp add: prems) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
332 |
thus "b \<le> d" by (simp add: preal_cancels) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
333 |
qed |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
334 |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
335 |
lemma real_le_lemma: |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
336 |
assumes l: "u1 + v2 \<le> u2 + v1" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
337 |
and "x1 + v1 = u1 + y1" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
338 |
and "x2 + v2 = u2 + y2" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
339 |
shows "x1 + y2 \<le> x2 + (y1::preal)" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
340 |
proof - |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
341 |
have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
342 |
hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: preal_add_ac) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
343 |
also have "... \<le> (x2+y1) + (u2+v1)" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
344 |
by (simp add: prems preal_add_le_cancel_left) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
345 |
finally show ?thesis by (simp add: preal_add_le_cancel_right) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
346 |
qed |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
347 |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
348 |
lemma real_le: |
14484 | 349 |
"(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) = |
350 |
(x1 + y2 \<le> x2 + y1)" |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
351 |
apply (simp add: real_le_def) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
352 |
apply (auto intro: real_le_lemma) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
353 |
done |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
354 |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
355 |
lemma real_le_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)" |
15542 | 356 |
by (cases z, cases w, simp add: real_le) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
357 |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
358 |
lemma real_trans_lemma: |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
359 |
assumes "x + v \<le> u + y" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
360 |
and "u + v' \<le> u' + v" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
361 |
and "x2 + v2 = u2 + y2" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
362 |
shows "x + v' \<le> u' + (y::preal)" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
363 |
proof - |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
364 |
have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: preal_add_ac) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
365 |
also have "... \<le> (u+y) + (u+v')" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
366 |
by (simp add: preal_add_le_cancel_right prems) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
367 |
also have "... \<le> (u+y) + (u'+v)" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
368 |
by (simp add: preal_add_le_cancel_left prems) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
369 |
also have "... = (u'+y) + (u+v)" by (simp add: preal_add_ac) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
370 |
finally show ?thesis by (simp add: preal_add_le_cancel_right) |
15542 | 371 |
qed |
14269 | 372 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
373 |
lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)" |
14484 | 374 |
apply (cases i, cases j, cases k) |
375 |
apply (simp add: real_le) |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
376 |
apply (blast intro: real_trans_lemma) |
14334 | 377 |
done |
378 |
||
379 |
(* Axiom 'order_less_le' of class 'order': *) |
|
380 |
lemma real_less_le: "((w::real) < z) = (w \<le> z & w \<noteq> z)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
381 |
by (simp add: real_less_def) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
382 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
383 |
instance real :: order |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
384 |
proof qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
385 |
(assumption | |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
386 |
rule real_le_refl real_le_trans real_le_anti_sym real_less_le)+ |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
387 |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
388 |
(* Axiom 'linorder_linear' of class 'linorder': *) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
389 |
lemma real_le_linear: "(z::real) \<le> w | w \<le> z" |
14484 | 390 |
apply (cases z, cases w) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
391 |
apply (auto simp add: real_le real_zero_def preal_add_ac preal_cancels) |
14334 | 392 |
done |
393 |
||
394 |
||
395 |
instance real :: linorder |
|
396 |
by (intro_classes, rule real_le_linear) |
|
397 |
||
398 |
||
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
399 |
lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))" |
14484 | 400 |
apply (cases x, cases y) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
401 |
apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
402 |
preal_add_ac) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
403 |
apply (simp_all add: preal_add_assoc [symmetric] preal_cancels) |
15542 | 404 |
done |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
405 |
|
14484 | 406 |
lemma real_add_left_mono: |
407 |
assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)" |
|
408 |
proof - |
|
409 |
have "z + x - (z + y) = (z + -z) + (x - y)" |
|
410 |
by (simp add: diff_minus add_ac) |
|
411 |
with le show ?thesis |
|
14754
a080eeeaec14
Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents:
14738
diff
changeset
|
412 |
by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus) |
14484 | 413 |
qed |
14334 | 414 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
415 |
lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
416 |
by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
417 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
418 |
lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
419 |
by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus) |
14334 | 420 |
|
421 |
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y" |
|
14484 | 422 |
apply (cases x, cases y) |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
423 |
apply (simp add: linorder_not_le [where 'a = real, symmetric] |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
424 |
linorder_not_le [where 'a = preal] |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
425 |
real_zero_def real_le real_mult) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
426 |
--{*Reduce to the (simpler) @{text "\<le>"} relation *} |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
427 |
apply (auto dest!: less_add_left_Ex |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
428 |
simp add: preal_add_ac preal_mult_ac |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
429 |
preal_add_mult_distrib2 preal_cancels preal_self_less_add_right) |
14334 | 430 |
done |
431 |
||
432 |
lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y" |
|
433 |
apply (rule real_sum_gt_zero_less) |
|
434 |
apply (drule real_less_sum_gt_zero [of x y]) |
|
435 |
apply (drule real_mult_order, assumption) |
|
436 |
apply (simp add: right_distrib) |
|
437 |
done |
|
438 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
439 |
text{*lemma for proving @{term "0<(1::real)"}*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
440 |
lemma real_zero_le_one: "0 \<le> (1::real)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
441 |
by (simp add: real_zero_def real_one_def real_le |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
442 |
preal_self_less_add_left order_less_imp_le) |
14334 | 443 |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
444 |
|
14334 | 445 |
subsection{*The Reals Form an Ordered Field*} |
446 |
||
447 |
instance real :: ordered_field |
|
448 |
proof |
|
449 |
fix x y z :: real |
|
450 |
show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono) |
|
451 |
show "x < y ==> 0 < z ==> z * x < z * y" by (simp add: real_mult_less_mono2) |
|
452 |
show "\<bar>x\<bar> = (if x < 0 then -x else x)" |
|
453 |
by (auto dest: order_le_less_trans simp add: real_abs_def linorder_not_le) |
|
454 |
qed |
|
455 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
456 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
457 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
458 |
text{*The function @{term real_of_preal} requires many proofs, but it seems |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
459 |
to be essential for proving completeness of the reals from that of the |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
460 |
positive reals.*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
461 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
462 |
lemma real_of_preal_add: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
463 |
"real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
464 |
by (simp add: real_of_preal_def real_add preal_add_mult_distrib preal_mult_1 |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
465 |
preal_add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
466 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
467 |
lemma real_of_preal_mult: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
468 |
"real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
469 |
by (simp add: real_of_preal_def real_mult preal_add_mult_distrib2 |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
470 |
preal_mult_1 preal_mult_1_right preal_add_ac preal_mult_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
471 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
472 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
473 |
text{*Gleason prop 9-4.4 p 127*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
474 |
lemma real_of_preal_trichotomy: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
475 |
"\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)" |
14484 | 476 |
apply (simp add: real_of_preal_def real_zero_def, cases x) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
477 |
apply (auto simp add: real_minus preal_add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
478 |
apply (cut_tac x = x and y = y in linorder_less_linear) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
479 |
apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc [symmetric]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
480 |
apply (auto simp add: preal_add_commute) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
481 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
482 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
483 |
lemma real_of_preal_leD: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
484 |
"real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2" |
14484 | 485 |
by (simp add: real_of_preal_def real_le preal_cancels) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
486 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
487 |
lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
488 |
by (auto simp add: real_of_preal_leD linorder_not_le [symmetric]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
489 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
490 |
lemma real_of_preal_lessD: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
491 |
"real_of_preal m1 < real_of_preal m2 ==> m1 < m2" |
14484 | 492 |
by (simp add: real_of_preal_def real_le linorder_not_le [symmetric] |
493 |
preal_cancels) |
|
494 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
495 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
496 |
lemma real_of_preal_less_iff [simp]: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
497 |
"(real_of_preal m1 < real_of_preal m2) = (m1 < m2)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
498 |
by (blast intro: real_of_preal_lessI real_of_preal_lessD) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
499 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
500 |
lemma real_of_preal_le_iff: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
501 |
"(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
502 |
by (simp add: linorder_not_less [symmetric]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
503 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
504 |
lemma real_of_preal_zero_less: "0 < real_of_preal m" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
505 |
apply (auto simp add: real_zero_def real_of_preal_def real_less_def real_le_def |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
506 |
preal_add_ac preal_cancels) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
507 |
apply (simp_all add: preal_add_assoc [symmetric] preal_cancels) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
508 |
apply (blast intro: preal_self_less_add_left order_less_imp_le) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
509 |
apply (insert preal_not_eq_self [of "preal_of_rat 1" m]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
510 |
apply (simp add: preal_add_ac) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
511 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
512 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
513 |
lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
514 |
by (simp add: real_of_preal_zero_less) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
515 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
516 |
lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m" |
14484 | 517 |
proof - |
518 |
from real_of_preal_minus_less_zero |
|
519 |
show ?thesis by (blast dest: order_less_trans) |
|
520 |
qed |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
521 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
522 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
523 |
subsection{*Theorems About the Ordering*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
524 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
525 |
text{*obsolete but used a lot*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
526 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
527 |
lemma real_not_refl2: "x < y ==> x \<noteq> (y::real)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
528 |
by blast |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
529 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
530 |
lemma real_le_imp_less_or_eq: "!!(x::real). x \<le> y ==> x < y | x = y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
531 |
by (simp add: order_le_less) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
532 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
533 |
lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
534 |
apply (auto simp add: real_of_preal_zero_less) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
535 |
apply (cut_tac x = x in real_of_preal_trichotomy) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
536 |
apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
537 |
done |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
538 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
539 |
lemma real_gt_preal_preal_Ex: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
540 |
"real_of_preal z < x ==> \<exists>y. x = real_of_preal y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
541 |
by (blast dest!: real_of_preal_zero_less [THEN order_less_trans] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
542 |
intro: real_gt_zero_preal_Ex [THEN iffD1]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
543 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
544 |
lemma real_ge_preal_preal_Ex: |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
545 |
"real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
546 |
by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
547 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
548 |
lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
549 |
by (auto elim: order_le_imp_less_or_eq [THEN disjE] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
550 |
intro: real_of_preal_zero_less [THEN [2] order_less_trans] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
551 |
simp add: real_of_preal_zero_less) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
552 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
553 |
lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
554 |
by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1]) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
555 |
|
14334 | 556 |
lemma real_add_less_le_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::real)" |
14738 | 557 |
by (rule OrderedGroup.add_less_le_mono) |
14334 | 558 |
|
559 |
lemma real_add_le_less_mono: |
|
560 |
"!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z" |
|
14738 | 561 |
by (rule OrderedGroup.add_le_less_mono) |
14334 | 562 |
|
563 |
lemma real_le_square [simp]: "(0::real) \<le> x*x" |
|
564 |
by (rule Ring_and_Field.zero_le_square) |
|
565 |
||
566 |
||
567 |
subsection{*More Lemmas*} |
|
568 |
||
569 |
lemma real_mult_left_cancel: "(c::real) \<noteq> 0 ==> (c*a=c*b) = (a=b)" |
|
570 |
by auto |
|
571 |
||
572 |
lemma real_mult_right_cancel: "(c::real) \<noteq> 0 ==> (a*c=b*c) = (a=b)" |
|
573 |
by auto |
|
574 |
||
575 |
text{*The precondition could be weakened to @{term "0\<le>x"}*} |
|
576 |
lemma real_mult_less_mono: |
|
577 |
"[| u<v; x<y; (0::real) < v; 0 < x |] ==> u*x < v* y" |
|
578 |
by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le) |
|
579 |
||
580 |
lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)" |
|
581 |
by (force elim: order_less_asym |
|
582 |
simp add: Ring_and_Field.mult_less_cancel_right) |
|
583 |
||
584 |
lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
585 |
apply (simp add: mult_le_cancel_right) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
586 |
apply (blast intro: elim: order_less_asym) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
587 |
done |
14334 | 588 |
|
589 |
lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)" |
|
15923 | 590 |
by(simp add:mult_commute) |
14334 | 591 |
|
592 |
text{*Only two uses?*} |
|
593 |
lemma real_mult_less_mono': |
|
594 |
"[| x < y; r1 < r2; (0::real) \<le> r1; 0 \<le> x|] ==> r1 * x < r2 * y" |
|
595 |
by (rule Ring_and_Field.mult_strict_mono') |
|
596 |
||
597 |
text{*FIXME: delete or at least combine the next two lemmas*} |
|
598 |
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)" |
|
14738 | 599 |
apply (drule OrderedGroup.equals_zero_I [THEN sym]) |
14334 | 600 |
apply (cut_tac x = y in real_le_square) |
14476 | 601 |
apply (auto, drule order_antisym, auto) |
14334 | 602 |
done |
603 |
||
604 |
lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)" |
|
605 |
apply (rule_tac y = x in real_sum_squares_cancel) |
|
14476 | 606 |
apply (simp add: add_commute) |
14334 | 607 |
done |
608 |
||
609 |
lemma real_add_order: "[| 0 < x; 0 < y |] ==> (0::real) < x + y" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
610 |
by (drule add_strict_mono [of concl: 0 0], assumption, simp) |
14334 | 611 |
|
612 |
lemma real_le_add_order: "[| 0 \<le> x; 0 \<le> y |] ==> (0::real) \<le> x + y" |
|
613 |
apply (drule order_le_imp_less_or_eq)+ |
|
614 |
apply (auto intro: real_add_order order_less_imp_le) |
|
615 |
done |
|
616 |
||
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
617 |
lemma real_inverse_unique: "x*y = (1::real) ==> y = inverse x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
618 |
apply (case_tac "x \<noteq> 0") |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
619 |
apply (rule_tac c1 = x in real_mult_left_cancel [THEN iffD1], auto) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
620 |
done |
14334 | 621 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
622 |
lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
623 |
by (auto dest: less_imp_inverse_less) |
14334 | 624 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
625 |
lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y" |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
626 |
proof - |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
627 |
have "0 + 0 \<le> x*x + y*y" by (blast intro: add_mono zero_le_square) |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
628 |
thus ?thesis by simp |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
629 |
qed |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
630 |
|
14334 | 631 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
632 |
subsection{*Embedding the Integers into the Reals*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
633 |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
634 |
defs (overloaded) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
635 |
real_of_nat_def: "real z == of_nat z" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
636 |
real_of_int_def: "real z == of_int z" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
637 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
638 |
lemma real_of_int_zero [simp]: "real (0::int) = 0" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
639 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
640 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
641 |
lemma real_of_one [simp]: "real (1::int) = (1::real)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
642 |
by (simp add: real_of_int_def) |
14334 | 643 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
644 |
lemma real_of_int_add: "real (x::int) + real y = real (x + y)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
645 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
646 |
declare real_of_int_add [symmetric, simp] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
647 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
648 |
lemma real_of_int_minus: "-real (x::int) = real (-x)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
649 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
650 |
declare real_of_int_minus [symmetric, simp] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
651 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
652 |
lemma real_of_int_diff: "real (x::int) - real y = real (x - y)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
653 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
654 |
declare real_of_int_diff [symmetric, simp] |
14334 | 655 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
656 |
lemma real_of_int_mult: "real (x::int) * real y = real (x * y)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
657 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
658 |
declare real_of_int_mult [symmetric, simp] |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
659 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
660 |
lemma real_of_int_zero_cancel [simp]: "(real x = 0) = (x = (0::int))" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
661 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
662 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
663 |
lemma real_of_int_inject [iff]: "(real (x::int) = real y) = (x = y)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
664 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
665 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
666 |
lemma real_of_int_less_iff [iff]: "(real (x::int) < real y) = (x < y)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
667 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
668 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
669 |
lemma real_of_int_le_iff [simp]: "(real (x::int) \<le> real y) = (x \<le> y)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
670 |
by (simp add: real_of_int_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
671 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
672 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
673 |
subsection{*Embedding the Naturals into the Reals*} |
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
674 |
|
14334 | 675 |
lemma real_of_nat_zero [simp]: "real (0::nat) = 0" |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
676 |
by (simp add: real_of_nat_def) |
14334 | 677 |
|
678 |
lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
679 |
by (simp add: real_of_nat_def) |
14334 | 680 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
681 |
lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
682 |
by (simp add: real_of_nat_def) |
14334 | 683 |
|
684 |
(*Not for addsimps: often the LHS is used to represent a positive natural*) |
|
685 |
lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)" |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
686 |
by (simp add: real_of_nat_def) |
14334 | 687 |
|
688 |
lemma real_of_nat_less_iff [iff]: |
|
689 |
"(real (n::nat) < real m) = (n < m)" |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
690 |
by (simp add: real_of_nat_def) |
14334 | 691 |
|
692 |
lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)" |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
693 |
by (simp add: real_of_nat_def) |
14334 | 694 |
|
695 |
lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)" |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
696 |
by (simp add: real_of_nat_def zero_le_imp_of_nat) |
14334 | 697 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
698 |
lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
699 |
by (simp add: real_of_nat_def del: of_nat_Suc) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
700 |
|
14334 | 701 |
lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
702 |
by (simp add: real_of_nat_def) |
14334 | 703 |
|
704 |
lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)" |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
705 |
by (simp add: real_of_nat_def) |
14334 | 706 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
707 |
lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
708 |
by (simp add: real_of_nat_def) |
14334 | 709 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
710 |
lemma real_of_nat_diff: "n \<le> m ==> real (m - n) = real (m::nat) - real n" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
711 |
by (simp add: add: real_of_nat_def) |
14334 | 712 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
713 |
lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
714 |
by (simp add: add: real_of_nat_def) |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
715 |
|
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
716 |
lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \<le> 0) = (n = 0)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
717 |
by (simp add: add: real_of_nat_def) |
14334 | 718 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
719 |
lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
720 |
by (simp add: add: real_of_nat_def) |
14334 | 721 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
722 |
lemma real_of_nat_ge_zero_cancel_iff [simp]: "(0 \<le> real (n::nat)) = (0 \<le> n)" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
723 |
by (simp add: add: real_of_nat_def) |
14334 | 724 |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14348
diff
changeset
|
725 |
lemma real_of_int_real_of_nat: "real (int n) = real n" |
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
726 |
by (simp add: real_of_nat_def real_of_int_def int_eq_of_nat) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14369
diff
changeset
|
727 |
|
14426 | 728 |
lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n" |
729 |
by (simp add: real_of_int_def real_of_nat_def) |
|
14334 | 730 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
731 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
732 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
733 |
subsection{*Numerals and Arithmetic*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
734 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
735 |
instance real :: number .. |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
736 |
|
15013 | 737 |
defs (overloaded) |
738 |
real_number_of_def: "(number_of w :: real) == of_int (Rep_Bin w)" |
|
739 |
--{*the type constraint is essential!*} |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
740 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
741 |
instance real :: number_ring |
15013 | 742 |
by (intro_classes, simp add: real_number_of_def) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
743 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
744 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
745 |
text{*Collapse applications of @{term real} to @{term number_of}*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
746 |
lemma real_number_of [simp]: "real (number_of v :: int) = number_of v" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
747 |
by (simp add: real_of_int_def of_int_number_of_eq) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
748 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
749 |
lemma real_of_nat_number_of [simp]: |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
750 |
"real (number_of v :: nat) = |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
751 |
(if neg (number_of v :: int) then 0 |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
752 |
else (number_of v :: real))" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
753 |
by (simp add: real_of_int_real_of_nat [symmetric] int_nat_number_of) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
754 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
755 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
756 |
use "real_arith.ML" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
757 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
758 |
setup real_arith_setup |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
759 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
760 |
subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
761 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
762 |
text{*Needed in this non-standard form by Hyperreal/Transcendental*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
763 |
lemma real_0_le_divide_iff: |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
764 |
"((0::real) \<le> x/y) = ((x \<le> 0 | 0 \<le> y) & (0 \<le> x | y \<le> 0))" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
765 |
by (simp add: real_divide_def zero_le_mult_iff, auto) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
766 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
767 |
lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
768 |
by arith |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
769 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15077
diff
changeset
|
770 |
lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
771 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
772 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15077
diff
changeset
|
773 |
lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
774 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
775 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15077
diff
changeset
|
776 |
lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
777 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
778 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15077
diff
changeset
|
779 |
lemma real_add_le_0_iff: "(x+y \<le> (0::real)) = (y \<le> -x)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
780 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
781 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15077
diff
changeset
|
782 |
lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
783 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
784 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
785 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
786 |
(* |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
787 |
FIXME: we should have this, as for type int, but many proofs would break. |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
788 |
It replaces x+-y by x-y. |
15086 | 789 |
declare real_diff_def [symmetric, simp] |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
790 |
*) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
791 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
792 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
793 |
subsubsection{*Density of the Reals*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
794 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
795 |
lemma real_lbound_gt_zero: |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
796 |
"[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
797 |
apply (rule_tac x = " (min d1 d2) /2" in exI) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
798 |
apply (simp add: min_def) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
799 |
done |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
800 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
801 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
802 |
text{*Similar results are proved in @{text Ring_and_Field}*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
803 |
lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
804 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
805 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
806 |
lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
807 |
by auto |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
808 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
809 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
810 |
subsection{*Absolute Value Function for the Reals*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
811 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
812 |
lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))" |
15003 | 813 |
by (simp add: abs_if) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
814 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
815 |
lemma abs_interval_iff: "(abs x < r) = (-r < x & x < (r::real))" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
816 |
by (force simp add: Ring_and_Field.abs_less_iff) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
817 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
818 |
lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))" |
14738 | 819 |
by (force simp add: OrderedGroup.abs_le_iff) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
820 |
|
14484 | 821 |
(*FIXME: used only once, in SEQ.ML*) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
822 |
lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)" |
15003 | 823 |
by (simp add: abs_if) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
824 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
825 |
lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)" |
15229 | 826 |
by (simp add: real_of_nat_ge_zero) |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
827 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
828 |
lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
829 |
apply (simp add: linorder_not_less) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
830 |
apply (auto intro: abs_ge_self [THEN order_trans]) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
831 |
done |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
832 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
833 |
text{*Used only in Hyperreal/Lim.ML*} |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
834 |
lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)" |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
835 |
apply (simp add: real_add_assoc) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
836 |
apply (rule_tac a1 = y in add_left_commute [THEN ssubst]) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
837 |
apply (rule real_add_assoc [THEN subst]) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
838 |
apply (rule abs_triangle_ineq) |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
839 |
done |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
840 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
841 |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
842 |
|
14334 | 843 |
ML |
844 |
{* |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
845 |
val real_lbound_gt_zero = thm"real_lbound_gt_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
846 |
val real_less_half_sum = thm"real_less_half_sum"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
847 |
val real_gt_half_sum = thm"real_gt_half_sum"; |
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14335
diff
changeset
|
848 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
849 |
val abs_interval_iff = thm"abs_interval_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
850 |
val abs_le_interval_iff = thm"abs_le_interval_iff"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
851 |
val abs_add_one_gt_zero = thm"abs_add_one_gt_zero"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
852 |
val abs_add_one_not_less_self = thm"abs_add_one_not_less_self"; |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
853 |
val abs_sum_triangle_ineq = thm"abs_sum_triangle_ineq"; |
14334 | 854 |
*} |
10752
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents:
10648
diff
changeset
|
855 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
856 |
|
5588 | 857 |
end |