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