| author | paulson | 
| Tue, 01 Sep 1998 15:05:36 +0200 | |
| changeset 5418 | a895ab904b85 | 
| parent 5184 | 9b8547a9496a | 
| child 5459 | 1dbaf888f4e7 | 
| permissions | -rw-r--r-- | 
| 5078 | 1 | (* Title : Real.ML | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Description : The reals | |
| 5 | *) | |
| 6 | ||
| 7 | open Real; | |
| 8 | ||
| 9 | (*** Proving that realrel is an equivalence relation ***) | |
| 10 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 11 | Goal "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] \ | 
| 5078 | 12 | \ ==> x1 + y3 = x3 + y1"; | 
| 13 | by (res_inst_tac [("C","y2")] preal_add_right_cancel 1);
 | |
| 14 | by (rotate_tac 1 1 THEN dtac sym 1); | |
| 15 | by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 16 | by (rtac (preal_add_left_commute RS subst) 1); | |
| 17 | by (res_inst_tac [("x1","x1")] (preal_add_assoc RS subst) 1);
 | |
| 18 | by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 19 | qed "preal_trans_lemma"; | |
| 20 | ||
| 21 | (** Natural deduction for realrel **) | |
| 22 | ||
| 23 | Goalw [realrel_def] | |
| 24 | "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)"; | |
| 25 | by (Fast_tac 1); | |
| 26 | qed "realrel_iff"; | |
| 27 | ||
| 28 | Goalw [realrel_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 29 | "[| x1 + y2 = x2 + y1 |] ==> ((x1,y1),(x2,y2)): realrel"; | 
| 5078 | 30 | by (Fast_tac 1); | 
| 31 | qed "realrelI"; | |
| 32 | ||
| 33 | Goalw [realrel_def] | |
| 34 | "p: realrel --> (EX x1 y1 x2 y2. \ | |
| 35 | \ p = ((x1,y1),(x2,y2)) & x1 + y2 = x2 + y1)"; | |
| 36 | by (Fast_tac 1); | |
| 37 | qed "realrelE_lemma"; | |
| 38 | ||
| 39 | val [major,minor] = goal thy | |
| 40 | "[| p: realrel; \ | |
| 41 | \ !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2)); x1+y2 = x2+y1 \ | |
| 42 | \ |] ==> Q |] ==> Q"; | |
| 43 | by (cut_facts_tac [major RS (realrelE_lemma RS mp)] 1); | |
| 44 | by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); | |
| 45 | qed "realrelE"; | |
| 46 | ||
| 47 | AddSIs [realrelI]; | |
| 48 | AddSEs [realrelE]; | |
| 49 | ||
| 50 | Goal "(x,x): realrel"; | |
| 51 | by (stac surjective_pairing 1 THEN rtac (refl RS realrelI) 1); | |
| 52 | qed "realrel_refl"; | |
| 53 | ||
| 54 | Goalw [equiv_def, refl_def, sym_def, trans_def] | |
| 55 |     "equiv {x::(preal*preal).True} realrel";
 | |
| 56 | by (fast_tac (claset() addSIs [realrel_refl] | |
| 57 | addSEs [sym,preal_trans_lemma]) 1); | |
| 58 | qed "equiv_realrel"; | |
| 59 | ||
| 60 | val equiv_realrel_iff = | |
| 61 | [TrueI, TrueI] MRS | |
| 62 | ([CollectI, CollectI] MRS | |
| 63 | (equiv_realrel RS eq_equiv_class_iff)); | |
| 64 | ||
| 65 | Goalw  [real_def,realrel_def,quotient_def] "realrel^^{(x,y)}:real";
 | |
| 66 | by (Blast_tac 1); | |
| 67 | qed "realrel_in_real"; | |
| 68 | ||
| 69 | Goal "inj_on Abs_real real"; | |
| 70 | by (rtac inj_on_inverseI 1); | |
| 71 | by (etac Abs_real_inverse 1); | |
| 72 | qed "inj_on_Abs_real"; | |
| 73 | ||
| 74 | Addsimps [equiv_realrel_iff,inj_on_Abs_real RS inj_on_iff, | |
| 75 | realrel_iff, realrel_in_real, Abs_real_inverse]; | |
| 76 | ||
| 77 | Addsimps [equiv_realrel RS eq_equiv_class_iff]; | |
| 78 | val eq_realrelD = equiv_realrel RSN (2,eq_equiv_class); | |
| 79 | ||
| 80 | Goal "inj(Rep_real)"; | |
| 81 | by (rtac inj_inverseI 1); | |
| 82 | by (rtac Rep_real_inverse 1); | |
| 83 | qed "inj_Rep_real"; | |
| 84 | ||
| 85 | (** real_preal: the injection from preal to real **) | |
| 86 | Goal "inj(real_preal)"; | |
| 87 | by (rtac injI 1); | |
| 88 | by (rewtac real_preal_def); | |
| 89 | by (dtac (inj_on_Abs_real RS inj_onD) 1); | |
| 90 | by (REPEAT (rtac realrel_in_real 1)); | |
| 91 | by (dtac eq_equiv_class 1); | |
| 92 | by (rtac equiv_realrel 1); | |
| 93 | by (Fast_tac 1); | |
| 94 | by Safe_tac; | |
| 95 | by (Asm_full_simp_tac 1); | |
| 96 | qed "inj_real_preal"; | |
| 97 | ||
| 98 | val [prem] = goal thy | |
| 99 |     "(!!x y. z = Abs_real(realrel^^{(x,y)}) ==> P) ==> P";
 | |
| 100 | by (res_inst_tac [("x1","z")] 
 | |
| 101 | (rewrite_rule [real_def] Rep_real RS quotientE) 1); | |
| 102 | by (dres_inst_tac [("f","Abs_real")] arg_cong 1);
 | |
| 103 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 104 | by (rtac prem 1); | |
| 105 | by (asm_full_simp_tac (simpset() addsimps [Rep_real_inverse]) 1); | |
| 106 | qed "eq_Abs_real"; | |
| 107 | ||
| 108 | (**** real_minus: additive inverse on real ****) | |
| 109 | ||
| 110 | Goalw [congruent_def] | |
| 111 |   "congruent realrel (%p. split (%x y. realrel^^{(y,x)}) p)";
 | |
| 112 | by Safe_tac; | |
| 113 | by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1); | |
| 114 | qed "real_minus_congruent"; | |
| 115 | ||
| 116 | (*Resolve th against the corresponding facts for real_minus*) | |
| 117 | val real_minus_ize = RSLIST [equiv_realrel, real_minus_congruent]; | |
| 118 | ||
| 119 | Goalw [real_minus_def] | |
| 120 |       "%~ (Abs_real(realrel^^{(x,y)})) = Abs_real(realrel ^^ {(y,x)})";
 | |
| 121 | by (res_inst_tac [("f","Abs_real")] arg_cong 1);
 | |
| 122 | by (simp_tac (simpset() addsimps | |
| 123 | [realrel_in_real RS Abs_real_inverse,real_minus_ize UN_equiv_class]) 1); | |
| 124 | qed "real_minus"; | |
| 125 | ||
| 126 | Goal "%~ (%~ z) = z"; | |
| 127 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 128 | by (asm_simp_tac (simpset() addsimps [real_minus]) 1); | |
| 129 | qed "real_minus_minus"; | |
| 130 | ||
| 131 | Addsimps [real_minus_minus]; | |
| 132 | ||
| 133 | Goal "inj(real_minus)"; | |
| 134 | by (rtac injI 1); | |
| 135 | by (dres_inst_tac [("f","real_minus")] arg_cong 1);
 | |
| 136 | by (asm_full_simp_tac (simpset() addsimps [real_minus_minus]) 1); | |
| 137 | qed "inj_real_minus"; | |
| 138 | ||
| 139 | Goalw [real_zero_def] "%~0r = 0r"; | |
| 140 | by (simp_tac (simpset() addsimps [real_minus]) 1); | |
| 141 | qed "real_minus_zero"; | |
| 142 | ||
| 143 | Addsimps [real_minus_zero]; | |
| 144 | ||
| 145 | Goal "(%~x = 0r) = (x = 0r)"; | |
| 146 | by (res_inst_tac [("z","x")] eq_Abs_real 1);
 | |
| 147 | by (auto_tac (claset(),simpset() addsimps [real_zero_def, | |
| 148 | real_minus] @ preal_add_ac)); | |
| 149 | qed "real_minus_zero_iff"; | |
| 150 | ||
| 151 | Addsimps [real_minus_zero_iff]; | |
| 152 | ||
| 153 | Goal "(%~x ~= 0r) = (x ~= 0r)"; | |
| 154 | by Auto_tac; | |
| 155 | qed "real_minus_not_zero_iff"; | |
| 156 | ||
| 157 | (*** Congruence property for addition ***) | |
| 158 | Goalw [congruent2_def] | |
| 159 | "congruent2 realrel (%p1 p2. \ | |
| 160 | \         split (%x1 y1. split (%x2 y2. realrel^^{(x1+x2, y1+y2)}) p2) p1)";
 | |
| 161 | by Safe_tac; | |
| 162 | by (asm_simp_tac (simpset() addsimps [preal_add_assoc]) 1); | |
| 163 | by (res_inst_tac [("z1.1","x1a")] (preal_add_left_commute RS ssubst) 1);
 | |
| 164 | by (asm_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1); | |
| 165 | by (asm_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 166 | qed "real_add_congruent2"; | |
| 167 | ||
| 168 | (*Resolve th against the corresponding facts for real_add*) | |
| 169 | val real_add_ize = RSLIST [equiv_realrel, real_add_congruent2]; | |
| 170 | ||
| 171 | Goalw [real_add_def] | |
| 172 |   "Abs_real(realrel^^{(x1,y1)}) + Abs_real(realrel^^{(x2,y2)}) = \
 | |
| 173 | \  Abs_real(realrel^^{(x1+x2, y1+y2)})";
 | |
| 174 | by (asm_simp_tac | |
| 175 | (simpset() addsimps [real_add_ize UN_equiv_class2]) 1); | |
| 176 | qed "real_add"; | |
| 177 | ||
| 178 | Goal "(z::real) + w = w + z"; | |
| 179 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 180 | by (res_inst_tac [("z","w")] eq_Abs_real 1);
 | |
| 181 | by (asm_simp_tac (simpset() addsimps (preal_add_ac @ [real_add])) 1); | |
| 182 | qed "real_add_commute"; | |
| 183 | ||
| 184 | Goal "((z1::real) + z2) + z3 = z1 + (z2 + z3)"; | |
| 185 | by (res_inst_tac [("z","z1")] eq_Abs_real 1);
 | |
| 186 | by (res_inst_tac [("z","z2")] eq_Abs_real 1);
 | |
| 187 | by (res_inst_tac [("z","z3")] eq_Abs_real 1);
 | |
| 188 | by (asm_simp_tac (simpset() addsimps [real_add, preal_add_assoc]) 1); | |
| 189 | qed "real_add_assoc"; | |
| 190 | ||
| 191 | (*For AC rewriting*) | |
| 192 | Goal "(x::real)+(y+z)=y+(x+z)"; | |
| 193 | by (rtac (real_add_commute RS trans) 1); | |
| 194 | by (rtac (real_add_assoc RS trans) 1); | |
| 195 | by (rtac (real_add_commute RS arg_cong) 1); | |
| 196 | qed "real_add_left_commute"; | |
| 197 | ||
| 198 | (* real addition is an AC operator *) | |
| 199 | val real_add_ac = [real_add_assoc,real_add_commute,real_add_left_commute]; | |
| 200 | ||
| 201 | Goalw [real_preal_def,real_zero_def] "0r + z = z"; | |
| 202 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 203 | by (asm_full_simp_tac (simpset() addsimps [real_add] @ preal_add_ac) 1); | |
| 204 | qed "real_add_zero_left"; | |
| 205 | ||
| 206 | Goal "z + 0r = z"; | |
| 207 | by (simp_tac (simpset() addsimps [real_add_zero_left,real_add_commute]) 1); | |
| 208 | qed "real_add_zero_right"; | |
| 209 | ||
| 210 | Goalw [real_zero_def] "z + %~z = 0r"; | |
| 211 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 212 | by (asm_full_simp_tac (simpset() addsimps [real_minus, | |
| 213 | real_add, preal_add_commute]) 1); | |
| 214 | qed "real_add_minus"; | |
| 215 | ||
| 216 | Goal "%~z + z = 0r"; | |
| 217 | by (simp_tac (simpset() addsimps | |
| 218 | [real_add_commute,real_add_minus]) 1); | |
| 219 | qed "real_add_minus_left"; | |
| 220 | ||
| 221 | Goal "? y. (x::real) + y = 0r"; | |
| 222 | by (fast_tac (claset() addIs [real_add_minus]) 1); | |
| 223 | qed "real_minus_ex"; | |
| 224 | ||
| 225 | Goal "?! y. (x::real) + y = 0r"; | |
| 226 | by (auto_tac (claset() addIs [real_add_minus],simpset())); | |
| 227 | by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
 | |
| 228 | by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1); | |
| 229 | by (asm_full_simp_tac (simpset() addsimps [real_add_commute, | |
| 230 | real_add_zero_right,real_add_zero_left]) 1); | |
| 231 | qed "real_minus_ex1"; | |
| 232 | ||
| 233 | Goal "?! y. y + (x::real) = 0r"; | |
| 234 | by (auto_tac (claset() addIs [real_add_minus_left],simpset())); | |
| 235 | by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
 | |
| 236 | by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1); | |
| 237 | by (asm_full_simp_tac (simpset() addsimps [real_add_commute, | |
| 238 | real_add_zero_right,real_add_zero_left]) 1); | |
| 239 | qed "real_minus_left_ex1"; | |
| 240 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 241 | Goal "x + y = 0r ==> x = %~y"; | 
| 5078 | 242 | by (cut_inst_tac [("z","y")] real_add_minus_left 1);
 | 
| 243 | by (res_inst_tac [("x1","y")] (real_minus_left_ex1 RS ex1E) 1);
 | |
| 244 | by (Blast_tac 1); | |
| 245 | qed "real_add_minus_eq_minus"; | |
| 246 | ||
| 247 | Goal "? y. x = %~y"; | |
| 248 | by (cut_inst_tac [("x","x")] real_minus_ex 1);
 | |
| 249 | by (etac exE 1 THEN dtac real_add_minus_eq_minus 1); | |
| 250 | by (Fast_tac 1); | |
| 251 | qed "real_as_add_inverse_ex"; | |
| 252 | ||
| 253 | (* real_minus_add_distrib *) | |
| 254 | Goal "%~(x + y) = %~x + %~y"; | |
| 255 | by (res_inst_tac [("z","x")] eq_Abs_real 1);
 | |
| 256 | by (res_inst_tac [("z","y")] eq_Abs_real 1);
 | |
| 257 | by (auto_tac (claset(),simpset() addsimps [real_minus,real_add])); | |
| 258 | qed "real_minus_add_eq"; | |
| 259 | ||
| 260 | val real_minus_add_distrib = real_minus_add_eq; | |
| 261 | ||
| 262 | Goal "((x::real) + y = x + z) = (y = z)"; | |
| 263 | by (Step_tac 1); | |
| 264 | by (dres_inst_tac [("f","%t.%~x + t")] arg_cong 1);
 | |
| 265 | by (asm_full_simp_tac (simpset() addsimps [real_add_minus_left, | |
| 266 | real_add_assoc RS sym,real_add_zero_left]) 1); | |
| 267 | qed "real_add_left_cancel"; | |
| 268 | ||
| 269 | Goal "(y + (x::real)= z + x) = (y = z)"; | |
| 270 | by (simp_tac (simpset() addsimps [real_add_commute,real_add_left_cancel]) 1); | |
| 271 | qed "real_add_right_cancel"; | |
| 272 | ||
| 273 | (*** Congruence property for multiplication ***) | |
| 274 | Goal "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> \ | |
| 275 | \ x * x1 + y * y1 + (x * y2 + x2 * y) = \ | |
| 276 | \ x * x2 + y * y2 + (x * y1 + x1 * y)"; | |
| 277 | by (asm_full_simp_tac (simpset() addsimps [preal_add_left_commute, | |
| 278 | preal_add_assoc RS sym,preal_add_mult_distrib2 RS sym]) 1); | |
| 279 | by (rtac (preal_mult_commute RS subst) 1); | |
| 280 | by (res_inst_tac [("y1","x2")] (preal_mult_commute RS subst) 1);
 | |
| 281 | by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc, | |
| 282 | preal_add_mult_distrib2 RS sym]) 1); | |
| 283 | by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1); | |
| 284 | qed "real_mult_congruent2_lemma"; | |
| 285 | ||
| 286 | Goal | |
| 287 | "congruent2 realrel (%p1 p2. \ | |
| 288 | \         split (%x1 y1. split (%x2 y2. realrel^^{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)";
 | |
| 289 | by (rtac (equiv_realrel RS congruent2_commuteI) 1); | |
| 290 | by Safe_tac; | |
| 291 | by (rewtac split_def); | |
| 292 | by (asm_simp_tac (simpset() addsimps [preal_mult_commute,preal_add_commute]) 1); | |
| 293 | by (auto_tac (claset(),simpset() addsimps [real_mult_congruent2_lemma])); | |
| 294 | qed "real_mult_congruent2"; | |
| 295 | ||
| 296 | (*Resolve th against the corresponding facts for real_mult*) | |
| 297 | val real_mult_ize = RSLIST [equiv_realrel, real_mult_congruent2]; | |
| 298 | ||
| 299 | Goalw [real_mult_def] | |
| 300 |    "Abs_real((realrel^^{(x1,y1)})) * Abs_real((realrel^^{(x2,y2)})) =   \
 | |
| 301 | \   Abs_real(realrel ^^ {(x1*x2+y1*y2,x1*y2+x2*y1)})";
 | |
| 302 | by (simp_tac (simpset() addsimps [real_mult_ize UN_equiv_class2]) 1); | |
| 303 | qed "real_mult"; | |
| 304 | ||
| 305 | Goal "(z::real) * w = w * z"; | |
| 306 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 307 | by (res_inst_tac [("z","w")] eq_Abs_real 1);
 | |
| 308 | by (asm_simp_tac (simpset() addsimps ([real_mult] @ preal_add_ac @ preal_mult_ac)) 1); | |
| 309 | qed "real_mult_commute"; | |
| 310 | ||
| 311 | Goal "((z1::real) * z2) * z3 = z1 * (z2 * z3)"; | |
| 312 | by (res_inst_tac [("z","z1")] eq_Abs_real 1);
 | |
| 313 | by (res_inst_tac [("z","z2")] eq_Abs_real 1);
 | |
| 314 | by (res_inst_tac [("z","z3")] eq_Abs_real 1);
 | |
| 315 | by (asm_simp_tac (simpset() addsimps ([preal_add_mult_distrib2,real_mult] @ | |
| 316 | preal_add_ac @ preal_mult_ac)) 1); | |
| 317 | qed "real_mult_assoc"; | |
| 318 | ||
| 319 | qed_goal "real_mult_left_commute" thy | |
| 320 | "(z1::real) * (z2 * z3) = z2 * (z1 * z3)" | |
| 321 | (fn _ => [rtac (real_mult_commute RS trans) 1, rtac (real_mult_assoc RS trans) 1, | |
| 322 | rtac (real_mult_commute RS arg_cong) 1]); | |
| 323 | ||
| 324 | (* real multiplication is an AC operator *) | |
| 325 | val real_mult_ac = [real_mult_assoc, real_mult_commute, real_mult_left_commute]; | |
| 326 | ||
| 327 | Goalw [real_one_def,pnat_one_def] "1r * z = z"; | |
| 328 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 329 | by (asm_full_simp_tac (simpset() addsimps [real_mult, | |
| 330 | preal_add_mult_distrib2,preal_mult_1_right] | |
| 331 | @ preal_mult_ac @ preal_add_ac) 1); | |
| 332 | qed "real_mult_1"; | |
| 333 | ||
| 334 | Goal "z * 1r = z"; | |
| 335 | by (simp_tac (simpset() addsimps [real_mult_commute, | |
| 336 | real_mult_1]) 1); | |
| 337 | qed "real_mult_1_right"; | |
| 338 | ||
| 339 | Goalw [real_zero_def,pnat_one_def] "0r * z = 0r"; | |
| 340 | by (res_inst_tac [("z","z")] eq_Abs_real 1);
 | |
| 341 | by (asm_full_simp_tac (simpset() addsimps [real_mult, | |
| 342 | preal_add_mult_distrib2,preal_mult_1_right] | |
| 343 | @ preal_mult_ac @ preal_add_ac) 1); | |
| 344 | qed "real_mult_0"; | |
| 345 | ||
| 346 | Goal "z * 0r = 0r"; | |
| 347 | by (simp_tac (simpset() addsimps [real_mult_commute, | |
| 348 | real_mult_0]) 1); | |
| 349 | qed "real_mult_0_right"; | |
| 350 | ||
| 351 | Addsimps [real_mult_0_right,real_mult_0]; | |
| 352 | ||
| 353 | Goal "%~(x * y) = %~x * y"; | |
| 354 | by (res_inst_tac [("z","x")] eq_Abs_real 1);
 | |
| 355 | by (res_inst_tac [("z","y")] eq_Abs_real 1);
 | |
| 356 | by (auto_tac (claset(),simpset() addsimps [real_minus,real_mult] | |
| 357 | @ preal_mult_ac @ preal_add_ac)); | |
| 358 | qed "real_minus_mult_eq1"; | |
| 359 | ||
| 360 | Goal "%~(x * y) = x * %~y"; | |
| 361 | by (res_inst_tac [("z","x")] eq_Abs_real 1);
 | |
| 362 | by (res_inst_tac [("z","y")] eq_Abs_real 1);
 | |
| 363 | by (auto_tac (claset(),simpset() addsimps [real_minus,real_mult] | |
| 364 | @ preal_mult_ac @ preal_add_ac)); | |
| 365 | qed "real_minus_mult_eq2"; | |
| 366 | ||
| 367 | Goal "%~x*%~y = x*y"; | |
| 368 | by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym, | |
| 369 | real_minus_mult_eq1 RS sym]) 1); | |
| 370 | qed "real_minus_mult_cancel"; | |
| 371 | ||
| 372 | Addsimps [real_minus_mult_cancel]; | |
| 373 | ||
| 374 | Goal "%~x*y = x*%~y"; | |
| 375 | by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym, | |
| 376 | real_minus_mult_eq1 RS sym]) 1); | |
| 377 | qed "real_minus_mult_commute"; | |
| 378 | ||
| 379 | (*----------------------------------------------------------------------------- | |
| 380 | ||
| 381 | -----------------------------------------------------------------------------*) | |
| 382 | ||
| 383 | (** Lemmas **) | |
| 384 | ||
| 385 | qed_goal "real_add_assoc_cong" thy | |
| 386 | "!!z. (z::real) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" | |
| 387 | (fn _ => [(asm_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1)]); | |
| 388 | ||
| 389 | qed_goal "real_add_assoc_swap" thy "(z::real) + (v + w) = v + (z + w)" | |
| 390 | (fn _ => [(REPEAT (ares_tac [real_add_commute RS real_add_assoc_cong] 1))]); | |
| 391 | ||
| 392 | Goal "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"; | |
| 393 | by (res_inst_tac [("z","z1")] eq_Abs_real 1);
 | |
| 394 | by (res_inst_tac [("z","z2")] eq_Abs_real 1);
 | |
| 395 | by (res_inst_tac [("z","w")] eq_Abs_real 1);
 | |
| 396 | by (asm_simp_tac | |
| 397 | (simpset() addsimps ([preal_add_mult_distrib2, real_add, real_mult] @ | |
| 398 | preal_add_ac @ preal_mult_ac)) 1); | |
| 399 | qed "real_add_mult_distrib"; | |
| 400 | ||
| 401 | val real_mult_commute'= read_instantiate [("z","w")] real_mult_commute;
 | |
| 402 | ||
| 403 | Goal "(w::real) * (z1 + z2) = (w * z1) + (w * z2)"; | |
| 404 | by (simp_tac (simpset() addsimps [real_mult_commute',real_add_mult_distrib]) 1); | |
| 405 | qed "real_add_mult_distrib2"; | |
| 406 | ||
| 407 | val real_mult_simps = [real_mult_1, real_mult_1_right]; | |
| 408 | Addsimps real_mult_simps; | |
| 409 | ||
| 410 | (*** one and zero are distinct ***) | |
| 411 | Goalw [real_zero_def,real_one_def] "0r ~= 1r"; | |
| 412 | by (auto_tac (claset(),simpset() addsimps | |
| 413 | [preal_self_less_add_left RS preal_not_refl2])); | |
| 414 | qed "real_zero_not_eq_one"; | |
| 415 | ||
| 416 | (*** existence of inverse ***) | |
| 417 | (** lemma -- alternative definition for 0r **) | |
| 418 | Goalw [real_zero_def] "0r = Abs_real (realrel ^^ {(x, x)})";
 | |
| 419 | by (auto_tac (claset(),simpset() addsimps [preal_add_commute])); | |
| 420 | qed "real_zero_iff"; | |
| 421 | ||
| 422 | Goalw [real_zero_def,real_one_def] | |
| 423 | "!!(x::real). x ~= 0r ==> ? y. x*y = 1r"; | |
| 424 | by (res_inst_tac [("z","x")] eq_Abs_real 1);
 | |
| 425 | by (cut_inst_tac [("r1.0","xa"),("r2.0","y")] preal_linear 1);
 | |
| 426 | by (auto_tac (claset() addSDs [preal_less_add_left_Ex], | |
| 427 | simpset() addsimps [real_zero_iff RS sym])); | |
| 428 | by (res_inst_tac [("x","Abs_real (realrel ^^ {(@#$#1p,pinv(D)+@#$#1p)})")] exI 1);
 | |
| 429 | by (res_inst_tac [("x","Abs_real (realrel ^^ {(pinv(D)+@#$#1p,@#$#1p)})")] exI 2);
 | |
| 430 | by (auto_tac (claset(),simpset() addsimps [real_mult, | |
| 431 | pnat_one_def,preal_mult_1_right,preal_add_mult_distrib2, | |
| 432 | preal_add_mult_distrib,preal_mult_1,preal_mult_inv_right] | |
| 433 | @ preal_add_ac @ preal_mult_ac)); | |
| 434 | qed "real_mult_inv_right_ex"; | |
| 435 | ||
| 436 | Goal "!!(x::real). x ~= 0r ==> ? y. y*x = 1r"; | |
| 437 | by (asm_simp_tac (simpset() addsimps [real_mult_commute, | |
| 438 | real_mult_inv_right_ex]) 1); | |
| 439 | qed "real_mult_inv_left_ex"; | |
| 440 | ||
| 441 | Goalw [rinv_def] "!!(x::real). x ~= 0r ==> rinv(x)*x = 1r"; | |
| 442 | by (forward_tac [real_mult_inv_left_ex] 1); | |
| 443 | by (Step_tac 1); | |
| 444 | by (rtac selectI2 1); | |
| 445 | by Auto_tac; | |
| 446 | qed "real_mult_inv_left"; | |
| 447 | ||
| 448 | Goal "!!(x::real). x ~= 0r ==> x*rinv(x) = 1r"; | |
| 449 | by (auto_tac (claset() addIs [real_mult_commute RS subst], | |
| 450 | simpset() addsimps [real_mult_inv_left])); | |
| 451 | qed "real_mult_inv_right"; | |
| 452 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 453 | Goal "(c::real) ~= 0r ==> (c*a=c*b) = (a=b)"; | 
| 5078 | 454 | by Auto_tac; | 
| 455 | by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
 | |
| 456 | by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac) 1); | |
| 457 | qed "real_mult_left_cancel"; | |
| 458 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 459 | Goal "(c::real) ~= 0r ==> (a*c=b*c) = (a=b)"; | 
| 5078 | 460 | by (Step_tac 1); | 
| 461 | by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
 | |
| 462 | by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac) 1); | |
| 463 | qed "real_mult_right_cancel"; | |
| 464 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 465 | Goalw [rinv_def] "x ~= 0r ==> rinv(x) ~= 0r"; | 
| 5078 | 466 | by (forward_tac [real_mult_inv_left_ex] 1); | 
| 467 | by (etac exE 1); | |
| 468 | by (rtac selectI2 1); | |
| 469 | by (auto_tac (claset(),simpset() addsimps [real_mult_0, | |
| 470 | real_zero_not_eq_one])); | |
| 471 | qed "rinv_not_zero"; | |
| 472 | ||
| 473 | Addsimps [real_mult_inv_left,real_mult_inv_right]; | |
| 474 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 475 | Goal "x ~= 0r ==> rinv(rinv x) = x"; | 
| 5078 | 476 | by (res_inst_tac [("c1","rinv x")] (real_mult_right_cancel RS iffD1) 1);
 | 
| 477 | by (etac rinv_not_zero 1); | |
| 478 | by (auto_tac (claset() addDs [rinv_not_zero],simpset())); | |
| 479 | qed "real_rinv_rinv"; | |
| 480 | ||
| 481 | Goalw [rinv_def] "rinv(1r) = 1r"; | |
| 482 | by (cut_facts_tac [real_zero_not_eq_one RS | |
| 483 | not_sym RS real_mult_inv_left_ex] 1); | |
| 484 | by (etac exE 1); | |
| 485 | by (rtac selectI2 1); | |
| 486 | by (auto_tac (claset(),simpset() addsimps | |
| 487 | [real_zero_not_eq_one RS not_sym])); | |
| 488 | qed "real_rinv_1"; | |
| 489 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 490 | Goal "x ~= 0r ==> rinv(%~x) = %~rinv(x)"; | 
| 5078 | 491 | by (res_inst_tac [("c1","%~x")] (real_mult_right_cancel RS iffD1) 1);
 | 
| 492 | by Auto_tac; | |
| 493 | qed "real_minus_rinv"; | |
| 494 | ||
| 495 | (*** theorems for ordering ***) | |
| 496 | (* prove introduction and elimination rules for real_less *) | |
| 497 | ||
| 498 | Goalw [real_less_def] | |
| 499 | "P < (Q::real) = (EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & \ | |
| 500 | \ (x1,y1::preal):Rep_real(P) & \ | |
| 501 | \ (x2,y2):Rep_real(Q))"; | |
| 502 | by (Fast_tac 1); | |
| 503 | qed "real_less_iff"; | |
| 504 | ||
| 505 | Goalw [real_less_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 506 | "[| x1 + y2 < x2 + y1; (x1,y1::preal):Rep_real(P); \ | 
| 5078 | 507 | \ (x2,y2):Rep_real(Q) |] ==> P < (Q::real)"; | 
| 508 | by (Fast_tac 1); | |
| 509 | qed "real_lessI"; | |
| 510 | ||
| 511 | Goalw [real_less_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 512 | "!!P. [| R1 < (R2::real); \ | 
| 5078 | 513 | \ !!x1 x2 y1 y2. x1 + y2 < x2 + y1 ==> P; \ | 
| 514 | \ !!x1 y1. (x1,y1::preal):Rep_real(R1) ==> P; \ | |
| 515 | \ !!x2 y2. (x2,y2::preal):Rep_real(R2) ==> P |] \ | |
| 516 | \ ==> P"; | |
| 517 | by Auto_tac; | |
| 518 | qed "real_lessE"; | |
| 519 | ||
| 520 | Goalw [real_less_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 521 | "R1 < (R2::real) ==> (EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & \ | 
| 5078 | 522 | \ (x1,y1::preal):Rep_real(R1) & \ | 
| 523 | \ (x2,y2):Rep_real(R2))"; | |
| 524 | by (Fast_tac 1); | |
| 525 | qed "real_lessD"; | |
| 526 | ||
| 527 | (* real_less is a strong order i.e nonreflexive and transitive *) | |
| 528 | (*** lemmas ***) | |
| 529 | Goal "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y"; | |
| 530 | by (asm_simp_tac (simpset() addsimps [preal_add_commute]) 1); | |
| 531 | qed "preal_lemma_eq_rev_sum"; | |
| 532 | ||
| 533 | Goal "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1"; | |
| 534 | by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 535 | qed "preal_add_left_commute_cancel"; | |
| 536 | ||
| 537 | Goal | |
| 538 | "!!(x::preal). [| x + y2a = x2a + y; \ | |
| 539 | \ x + y2b = x2b + y |] \ | |
| 540 | \ ==> x2a + y2b = x2b + y2a"; | |
| 541 | by (dtac preal_lemma_eq_rev_sum 1); | |
| 542 | by (assume_tac 1); | |
| 543 | by (thin_tac "x + y2b = x2b + y" 1); | |
| 544 | by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 545 | by (dtac preal_add_left_commute_cancel 1); | |
| 546 | by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 547 | qed "preal_lemma_for_not_refl"; | |
| 548 | ||
| 549 | Goal "~ (R::real) < R"; | |
| 550 | by (res_inst_tac [("z","R")] eq_Abs_real 1);
 | |
| 551 | by (auto_tac (claset(),simpset() addsimps [real_less_def])); | |
| 552 | by (dtac preal_lemma_for_not_refl 1); | |
| 553 | by (assume_tac 1 THEN rotate_tac 2 1); | |
| 554 | by (auto_tac (claset(),simpset() addsimps [preal_less_not_refl])); | |
| 555 | qed "real_less_not_refl"; | |
| 556 | ||
| 557 | (*** y < y ==> P ***) | |
| 558 | bind_thm("real_less_irrefl",real_less_not_refl RS notE);
 | |
| 559 | ||
| 560 | Goal "!!(x::real). x < y ==> x ~= y"; | |
| 561 | by (auto_tac (claset(),simpset() addsimps [real_less_not_refl])); | |
| 562 | qed "real_not_refl2"; | |
| 563 | ||
| 564 | (* lemma re-arranging and eliminating terms *) | |
| 565 | Goal "!! (a::preal). [| a + b = c + d; \ | |
| 566 | \ x2b + d + (c + y2e) < a + y2b + (x2e + b) |] \ | |
| 567 | \ ==> x2b + y2e < x2e + y2b"; | |
| 568 | by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 569 | by (res_inst_tac [("C","c+d")] preal_add_left_less_cancel 1);
 | |
| 570 | by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1); | |
| 571 | qed "preal_lemma_trans"; | |
| 572 | ||
| 573 | (** heavy re-writing involved*) | |
| 574 | Goal "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3"; | |
| 575 | by (res_inst_tac [("z","R1")] eq_Abs_real 1);
 | |
| 576 | by (res_inst_tac [("z","R2")] eq_Abs_real 1);
 | |
| 577 | by (res_inst_tac [("z","R3")] eq_Abs_real 1);
 | |
| 578 | by (auto_tac (claset(),simpset() addsimps [real_less_def])); | |
| 579 | by (REPEAT(rtac exI 1)); | |
| 580 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 581 | by (REPEAT(Blast_tac 2)); | |
| 582 | by (dtac preal_lemma_for_not_refl 1 THEN assume_tac 1); | |
| 583 | by (blast_tac (claset() addDs [preal_add_less_mono] | |
| 584 | addIs [preal_lemma_trans]) 1); | |
| 585 | qed "real_less_trans"; | |
| 586 | ||
| 587 | Goal "!! (R1::real). [| R1 < R2; R2 < R1 |] ==> P"; | |
| 588 | by (dtac real_less_trans 1 THEN assume_tac 1); | |
| 589 | by (asm_full_simp_tac (simpset() addsimps [real_less_not_refl]) 1); | |
| 590 | qed "real_less_asym"; | |
| 591 | ||
| 592 | (****)(****)(****)(****)(****)(****)(****)(****)(****)(****) | |
| 593 | (****** Map and more real_less ******) | |
| 594 | (*** mapping from preal into real ***) | |
| 595 | Goalw [real_preal_def] | |
| 596 | "%#((z1::preal) + z2) = %#z1 + %#z2"; | |
| 597 | by (asm_simp_tac (simpset() addsimps [real_add, | |
| 598 | preal_add_mult_distrib,preal_mult_1] addsimps preal_add_ac) 1); | |
| 599 | qed "real_preal_add"; | |
| 600 | ||
| 601 | Goalw [real_preal_def] | |
| 602 | "%#((z1::preal) * z2) = %#z1* %#z2"; | |
| 603 | by (full_simp_tac (simpset() addsimps [real_mult, | |
| 604 | preal_add_mult_distrib2,preal_mult_1, | |
| 605 | preal_mult_1_right,pnat_one_def] | |
| 606 | @ preal_add_ac @ preal_mult_ac) 1); | |
| 607 | qed "real_preal_mult"; | |
| 608 | ||
| 609 | Goalw [real_preal_def] | |
| 610 |       "!!(x::preal). y < x ==> ? m. Abs_real (realrel ^^ {(x,y)}) = %#m";
 | |
| 611 | by (auto_tac (claset() addSDs [preal_less_add_left_Ex], | |
| 612 | simpset() addsimps preal_add_ac)); | |
| 613 | qed "real_preal_ExI"; | |
| 614 | ||
| 615 | Goalw [real_preal_def] | |
| 616 |       "!!(x::preal). ? m. Abs_real (realrel ^^ {(x,y)}) = %#m ==> y < x";
 | |
| 617 | by (auto_tac (claset(),simpset() addsimps | |
| 618 | [preal_add_commute,preal_add_assoc])); | |
| 619 | by (asm_full_simp_tac (simpset() addsimps | |
| 620 | [preal_add_assoc RS sym,preal_self_less_add_left]) 1); | |
| 621 | qed "real_preal_ExD"; | |
| 622 | ||
| 623 | Goal "(? m. Abs_real (realrel ^^ {(x,y)}) = %#m) = (y < x)";
 | |
| 624 | by (fast_tac (claset() addSIs [real_preal_ExI,real_preal_ExD]) 1); | |
| 625 | qed "real_preal_iff"; | |
| 626 | ||
| 627 | (*** Gleason prop 9-4.4 p 127 ***) | |
| 628 | Goalw [real_preal_def,real_zero_def] | |
| 629 | "? m. (x::real) = %#m | x = 0r | x = %~(%#m)"; | |
| 630 | by (res_inst_tac [("z","x")] eq_Abs_real 1);
 | |
| 631 | by (auto_tac (claset(),simpset() addsimps [real_minus] @ preal_add_ac)); | |
| 632 | by (cut_inst_tac [("r1.0","x"),("r2.0","y")] preal_linear 1);
 | |
| 633 | by (auto_tac (claset() addSDs [preal_less_add_left_Ex], | |
| 634 | simpset() addsimps [preal_add_assoc RS sym])); | |
| 635 | by (auto_tac (claset(),simpset() addsimps [preal_add_commute])); | |
| 636 | qed "real_preal_trichotomy"; | |
| 637 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 638 | Goal "!!P. [| !!m. x = %#m ==> P; \ | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 639 | \ x = 0r ==> P; \ | 
| 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 640 | \ !!m. x = %~(%#m) ==> P |] ==> P"; | 
| 5078 | 641 | by (cut_inst_tac [("x","x")] real_preal_trichotomy 1);
 | 
| 642 | by Auto_tac; | |
| 643 | qed "real_preal_trichotomyE"; | |
| 644 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 645 | Goalw [real_preal_def] "%#m1 < %#m2 ==> m1 < m2"; | 
| 5078 | 646 | by (auto_tac (claset(),simpset() addsimps [real_less_def] @ preal_add_ac)); | 
| 647 | by (auto_tac (claset(),simpset() addsimps [preal_add_assoc RS sym])); | |
| 648 | by (auto_tac (claset(),simpset() addsimps preal_add_ac)); | |
| 649 | qed "real_preal_lessD"; | |
| 650 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 651 | Goal "m1 < m2 ==> %#m1 < %#m2"; | 
| 5078 | 652 | by (dtac preal_less_add_left_Ex 1); | 
| 653 | by (auto_tac (claset(),simpset() addsimps [real_preal_add, | |
| 654 | real_preal_def,real_less_def])); | |
| 655 | by (REPEAT(rtac exI 1)); | |
| 656 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 657 | by (REPEAT(Fast_tac 2)); | |
| 658 | by (simp_tac (simpset() addsimps [preal_self_less_add_left] | |
| 659 | delsimps [preal_add_less_iff2]) 1); | |
| 660 | qed "real_preal_lessI"; | |
| 661 | ||
| 662 | Goal "(%#m1 < %#m2) = (m1 < m2)"; | |
| 663 | by (fast_tac (claset() addIs [real_preal_lessI,real_preal_lessD]) 1); | |
| 664 | qed "real_preal_less_iff1"; | |
| 665 | ||
| 666 | Addsimps [real_preal_less_iff1]; | |
| 667 | ||
| 668 | Goal "%~ %#m < %#m"; | |
| 669 | by (auto_tac (claset(),simpset() addsimps | |
| 670 | [real_preal_def,real_less_def,real_minus])); | |
| 671 | by (REPEAT(rtac exI 1)); | |
| 672 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 673 | by (REPEAT(Fast_tac 2)); | |
| 674 | by (full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 675 | by (full_simp_tac (simpset() addsimps [preal_self_less_add_right, | |
| 676 | preal_add_assoc RS sym]) 1); | |
| 677 | qed "real_preal_minus_less_self"; | |
| 678 | ||
| 679 | Goalw [real_zero_def] "%~ %#m < 0r"; | |
| 680 | by (auto_tac (claset(),simpset() addsimps | |
| 681 | [real_preal_def,real_less_def,real_minus])); | |
| 682 | by (REPEAT(rtac exI 1)); | |
| 683 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 684 | by (REPEAT(Fast_tac 2)); | |
| 685 | by (full_simp_tac (simpset() addsimps | |
| 686 | [preal_self_less_add_right] @ preal_add_ac) 1); | |
| 687 | qed "real_preal_minus_less_zero"; | |
| 688 | ||
| 689 | Goal "~ 0r < %~ %#m"; | |
| 690 | by (cut_facts_tac [real_preal_minus_less_zero] 1); | |
| 691 | by (fast_tac (claset() addDs [real_less_trans] | |
| 692 | addEs [real_less_irrefl]) 1); | |
| 693 | qed "real_preal_not_minus_gt_zero"; | |
| 694 | ||
| 695 | Goalw [real_zero_def] " 0r < %#m"; | |
| 696 | by (auto_tac (claset(),simpset() addsimps | |
| 697 | [real_preal_def,real_less_def,real_minus])); | |
| 698 | by (REPEAT(rtac exI 1)); | |
| 699 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 700 | by (REPEAT(Fast_tac 2)); | |
| 701 | by (full_simp_tac (simpset() addsimps | |
| 702 | [preal_self_less_add_right] @ preal_add_ac) 1); | |
| 703 | qed "real_preal_zero_less"; | |
| 704 | ||
| 705 | Goal "~ %#m < 0r"; | |
| 706 | by (cut_facts_tac [real_preal_zero_less] 1); | |
| 707 | by (fast_tac (claset() addDs [real_less_trans] | |
| 708 | addEs [real_less_irrefl]) 1); | |
| 709 | qed "real_preal_not_less_zero"; | |
| 710 | ||
| 711 | Goal "0r < %~ %~ %#m"; | |
| 712 | by (simp_tac (simpset() addsimps | |
| 713 | [real_preal_zero_less]) 1); | |
| 714 | qed "real_minus_minus_zero_less"; | |
| 715 | ||
| 716 | (* another lemma *) | |
| 717 | Goalw [real_zero_def] " 0r < %#m + %#m1"; | |
| 718 | by (auto_tac (claset(),simpset() addsimps | |
| 719 | [real_preal_def,real_less_def,real_add])); | |
| 720 | by (REPEAT(rtac exI 1)); | |
| 721 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 722 | by (REPEAT(Fast_tac 2)); | |
| 723 | by (full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 724 | by (full_simp_tac (simpset() addsimps [preal_self_less_add_right, | |
| 725 | preal_add_assoc RS sym]) 1); | |
| 726 | qed "real_preal_sum_zero_less"; | |
| 727 | ||
| 728 | Goal "%~ %#m < %#m1"; | |
| 729 | by (auto_tac (claset(),simpset() addsimps | |
| 730 | [real_preal_def,real_less_def,real_minus])); | |
| 731 | by (REPEAT(rtac exI 1)); | |
| 732 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 733 | by (REPEAT(Fast_tac 2)); | |
| 734 | by (full_simp_tac (simpset() addsimps preal_add_ac) 1); | |
| 735 | by (full_simp_tac (simpset() addsimps [preal_self_less_add_right, | |
| 736 | preal_add_assoc RS sym]) 1); | |
| 737 | qed "real_preal_minus_less_all"; | |
| 738 | ||
| 739 | Goal "~ %#m < %~ %#m1"; | |
| 740 | by (cut_facts_tac [real_preal_minus_less_all] 1); | |
| 741 | by (fast_tac (claset() addDs [real_less_trans] | |
| 742 | addEs [real_less_irrefl]) 1); | |
| 743 | qed "real_preal_not_minus_gt_all"; | |
| 744 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 745 | Goal "%~ %#m1 < %~ %#m2 ==> %#m2 < %#m1"; | 
| 5078 | 746 | by (auto_tac (claset(),simpset() addsimps | 
| 747 | [real_preal_def,real_less_def,real_minus])); | |
| 748 | by (REPEAT(rtac exI 1)); | |
| 749 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 750 | by (REPEAT(Fast_tac 2)); | |
| 751 | by (auto_tac (claset(),simpset() addsimps preal_add_ac)); | |
| 752 | by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1); | |
| 753 | by (auto_tac (claset(),simpset() addsimps preal_add_ac)); | |
| 754 | qed "real_preal_minus_less_rev1"; | |
| 755 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 756 | Goal "%#m1 < %#m2 ==> %~ %#m2 < %~ %#m1"; | 
| 5078 | 757 | by (auto_tac (claset(),simpset() addsimps | 
| 758 | [real_preal_def,real_less_def,real_minus])); | |
| 759 | by (REPEAT(rtac exI 1)); | |
| 760 | by (EVERY[rtac conjI 1, rtac conjI 2]); | |
| 761 | by (REPEAT(Fast_tac 2)); | |
| 762 | by (auto_tac (claset(),simpset() addsimps preal_add_ac)); | |
| 763 | by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1); | |
| 764 | by (auto_tac (claset(),simpset() addsimps preal_add_ac)); | |
| 765 | qed "real_preal_minus_less_rev2"; | |
| 766 | ||
| 767 | Goal "(%~ %#m1 < %~ %#m2) = (%#m2 < %#m1)"; | |
| 768 | by (blast_tac (claset() addSIs [real_preal_minus_less_rev1, | |
| 769 | real_preal_minus_less_rev2]) 1); | |
| 770 | qed "real_preal_minus_less_rev_iff"; | |
| 771 | ||
| 772 | Addsimps [real_preal_minus_less_rev_iff]; | |
| 773 | ||
| 774 | (*** linearity ***) | |
| 775 | Goal "(R1::real) < R2 | R1 = R2 | R2 < R1"; | |
| 776 | by (res_inst_tac [("x","R1")]  real_preal_trichotomyE 1);
 | |
| 777 | by (ALLGOALS(res_inst_tac [("x","R2")]  real_preal_trichotomyE));
 | |
| 778 | by (auto_tac (claset() addSDs [preal_le_anti_sym], | |
| 779 | simpset() addsimps [preal_less_le_iff,real_preal_minus_less_zero, | |
| 780 | real_preal_zero_less,real_preal_minus_less_all])); | |
| 781 | qed "real_linear"; | |
| 782 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 783 | Goal "!!(R1::real). [| R1 < R2 ==> P; R1 = R2 ==> P; \ | 
| 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 784 | \ R2 < R1 ==> P |] ==> P"; | 
| 5078 | 785 | by (cut_inst_tac [("R1.0","R1"),("R2.0","R2")] real_linear 1);
 | 
| 786 | by Auto_tac; | |
| 787 | qed "real_linear_less2"; | |
| 788 | ||
| 789 | (*** Properties of <= ***) | |
| 790 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 791 | Goalw [real_le_def] "~(w < z) ==> z <= (w::real)"; | 
| 5078 | 792 | by (assume_tac 1); | 
| 793 | qed "real_leI"; | |
| 794 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 795 | Goalw [real_le_def] "z<=w ==> ~(w<(z::real))"; | 
| 5078 | 796 | by (assume_tac 1); | 
| 797 | qed "real_leD"; | |
| 798 | ||
| 799 | val real_leE = make_elim real_leD; | |
| 800 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 801 | Goal "(~(w < z)) = (z <= (w::real))"; | 
| 5078 | 802 | by (fast_tac (claset() addSIs [real_leI,real_leD]) 1); | 
| 803 | qed "real_less_le_iff"; | |
| 804 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 805 | Goalw [real_le_def] "~ z <= w ==> w<(z::real)"; | 
| 5078 | 806 | by (Fast_tac 1); | 
| 807 | qed "not_real_leE"; | |
| 808 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 809 | Goalw [real_le_def] "z < w ==> z <= (w::real)"; | 
| 5078 | 810 | by (fast_tac (claset() addEs [real_less_asym]) 1); | 
| 811 | qed "real_less_imp_le"; | |
| 812 | ||
| 813 | Goalw [real_le_def] "!!(x::real). x <= y ==> x < y | x = y"; | |
| 814 | by (cut_facts_tac [real_linear] 1); | |
| 815 | by (fast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1); | |
| 816 | qed "real_le_imp_less_or_eq"; | |
| 817 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 818 | Goalw [real_le_def] "z<w | z=w ==> z <=(w::real)"; | 
| 5078 | 819 | by (cut_facts_tac [real_linear] 1); | 
| 820 | by (fast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1); | |
| 821 | qed "real_less_or_eq_imp_le"; | |
| 822 | ||
| 823 | Goal "(x <= (y::real)) = (x < y | x=y)"; | |
| 824 | by (REPEAT(ares_tac [iffI, real_less_or_eq_imp_le, real_le_imp_less_or_eq] 1)); | |
| 825 | qed "real_le_eq_less_or_eq"; | |
| 826 | ||
| 827 | Goal "w <= (w::real)"; | |
| 828 | by (simp_tac (simpset() addsimps [real_le_eq_less_or_eq]) 1); | |
| 829 | qed "real_le_refl"; | |
| 830 | ||
| 831 | val prems = goal Real.thy "!!i. [| i <= j; j < k |] ==> i < (k::real)"; | |
| 832 | by (dtac real_le_imp_less_or_eq 1); | |
| 833 | by (fast_tac (claset() addIs [real_less_trans]) 1); | |
| 834 | qed "real_le_less_trans"; | |
| 835 | ||
| 836 | Goal "!! (i::real). [| i < j; j <= k |] ==> i < k"; | |
| 837 | by (dtac real_le_imp_less_or_eq 1); | |
| 838 | by (fast_tac (claset() addIs [real_less_trans]) 1); | |
| 839 | qed "real_less_le_trans"; | |
| 840 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 841 | Goal "[| i <= j; j <= k |] ==> i <= (k::real)"; | 
| 5078 | 842 | by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq, | 
| 843 | rtac real_less_or_eq_imp_le, fast_tac (claset() addIs [real_less_trans])]); | |
| 844 | qed "real_le_trans"; | |
| 845 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 846 | Goal "[| z <= w; w <= z |] ==> z = (w::real)"; | 
| 5078 | 847 | by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq, | 
| 848 | fast_tac (claset() addEs [real_less_irrefl,real_less_asym])]); | |
| 849 | qed "real_le_anti_sym"; | |
| 850 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 851 | Goal "[| ~ y < x; y ~= x |] ==> x < (y::real)"; | 
| 5078 | 852 | by (rtac not_real_leE 1); | 
| 853 | by (fast_tac (claset() addDs [real_le_imp_less_or_eq]) 1); | |
| 854 | qed "not_less_not_eq_real_less"; | |
| 855 | ||
| 856 | Goal "(0r < %~R) = (R < 0r)"; | |
| 857 | by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
 | |
| 858 | by (auto_tac (claset(),simpset() addsimps [real_preal_not_minus_gt_zero, | |
| 859 | real_preal_not_less_zero,real_preal_zero_less, | |
| 860 | real_preal_minus_less_zero])); | |
| 861 | qed "real_minus_zero_less_iff"; | |
| 862 | ||
| 863 | Addsimps [real_minus_zero_less_iff]; | |
| 864 | ||
| 865 | Goal "(%~R < 0r) = (0r < R)"; | |
| 866 | by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
 | |
| 867 | by (auto_tac (claset(),simpset() addsimps [real_preal_not_minus_gt_zero, | |
| 868 | real_preal_not_less_zero,real_preal_zero_less, | |
| 869 | real_preal_minus_less_zero])); | |
| 870 | qed "real_minus_zero_less_iff2"; | |
| 871 | ||
| 872 | (** lemma **) | |
| 873 | Goal "(0r < x) = (? y. x = %#y)"; | |
| 874 | by (auto_tac (claset(),simpset() addsimps [real_preal_zero_less])); | |
| 875 | by (cut_inst_tac [("x","x")] real_preal_trichotomy 1);
 | |
| 876 | by (blast_tac (claset() addSEs [real_less_irrefl, | |
| 877 | real_preal_not_minus_gt_zero RS notE]) 1); | |
| 878 | qed "real_gt_zero_preal_Ex"; | |
| 879 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 880 | Goal "%#z < x ==> ? y. x = %#y"; | 
| 5078 | 881 | by (blast_tac (claset() addSDs [real_preal_zero_less RS real_less_trans] | 
| 882 | addIs [real_gt_zero_preal_Ex RS iffD1]) 1); | |
| 883 | qed "real_gt_preal_preal_Ex"; | |
| 884 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 885 | Goal "%#z <= x ==> ? y. x = %#y"; | 
| 5078 | 886 | by (blast_tac (claset() addDs [real_le_imp_less_or_eq, | 
| 887 | real_gt_preal_preal_Ex]) 1); | |
| 888 | qed "real_ge_preal_preal_Ex"; | |
| 889 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 890 | Goal "y <= 0r ==> !x. y < %#x"; | 
| 5078 | 891 | by (auto_tac (claset() addEs [real_le_imp_less_or_eq RS disjE] | 
| 892 | addIs [real_preal_zero_less RSN(2,real_less_trans)], | |
| 893 | simpset() addsimps [real_preal_zero_less])); | |
| 894 | qed "real_less_all_preal"; | |
| 895 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 896 | Goal "~ 0r < y ==> !x. y < %#x"; | 
| 5078 | 897 | by (blast_tac (claset() addSIs [real_less_all_preal,real_leI]) 1); | 
| 898 | qed "real_less_all_real2"; | |
| 899 | ||
| 900 | (**** Derive alternative definition for real_less ****) | |
| 901 | (** lemma **) | |
| 902 | Goal "!!(R::real). ? A. S + A = R"; | |
| 903 | by (res_inst_tac [("x","%~S + R")] exI 1);
 | |
| 904 | by (simp_tac (simpset() addsimps [real_add_minus, | |
| 905 | real_add_zero_right] @ real_add_ac) 1); | |
| 906 | qed "real_lemma_add_left_ex"; | |
| 907 | ||
| 908 | Goal "!!(R::real). R < S ==> ? T. R + T = S"; | |
| 909 | by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
 | |
| 910 | by (ALLGOALS(res_inst_tac [("x","S")]  real_preal_trichotomyE));
 | |
| 911 | by (auto_tac (claset() addSDs [preal_le_anti_sym] addSDs [preal_less_add_left_Ex], | |
| 912 | simpset() addsimps [preal_less_le_iff,real_preal_add,real_minus_add_eq, | |
| 913 | real_preal_minus_less_zero,real_less_not_refl,real_minus_ex,real_add_assoc, | |
| 914 | real_preal_zero_less,real_preal_minus_less_all,real_add_minus_left, | |
| 915 | real_preal_not_less_zero,real_add_zero_left,real_lemma_add_left_ex])); | |
| 916 | qed "real_less_add_left_Ex"; | |
| 917 | ||
| 918 | Goal "!!(R::real). R < S ==> ? T. 0r < T & R + T = S"; | |
| 919 | by (res_inst_tac [("x","R")]  real_preal_trichotomyE 1);
 | |
| 920 | by (ALLGOALS(res_inst_tac [("x","S")]  real_preal_trichotomyE));
 | |
| 921 | by (auto_tac (claset() addSDs [preal_less_add_left_Ex], | |
| 922 | simpset() addsimps [real_preal_not_minus_gt_all, | |
| 923 | real_preal_add, real_preal_not_less_zero,real_less_not_refl, | |
| 924 | real_preal_not_minus_gt_zero,real_add_zero_left,real_minus_add_eq])); | |
| 925 | by (res_inst_tac [("x","%#D")] exI 1);
 | |
| 926 | by (res_inst_tac [("x","%#m+%#ma")] exI 2);
 | |
| 927 | by (res_inst_tac [("x","%#m")] exI 3);
 | |
| 928 | by (res_inst_tac [("x","%#D")] exI 4);
 | |
| 929 | by (auto_tac (claset(),simpset() addsimps [real_preal_zero_less, | |
| 930 | real_preal_sum_zero_less,real_add_minus_left,real_add_assoc, | |
| 931 | real_add_minus,real_add_zero_right])); | |
| 932 | by (simp_tac (simpset() addsimps [real_add_assoc RS sym, | |
| 933 | real_add_minus_left,real_add_zero_left]) 1); | |
| 934 | qed "real_less_add_positive_left_Ex"; | |
| 935 | ||
| 936 | (* lemmas *) | |
| 937 | (** change naff name(s)! **) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 938 | Goal "(W < S) ==> (0r < S + %~W)"; | 
| 5078 | 939 | by (dtac real_less_add_positive_left_Ex 1); | 
| 940 | by (auto_tac (claset(),simpset() addsimps [real_add_minus, | |
| 941 | real_add_zero_right] @ real_add_ac)); | |
| 942 | qed "real_less_sum_gt_zero"; | |
| 943 | ||
| 944 | Goal "!!S. T = S + W ==> S = T + %~W"; | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 945 | by (asm_simp_tac (simpset() addsimps [real_add_minus, real_add_zero_right] | 
| 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 946 | @ real_add_ac) 1); | 
| 5078 | 947 | qed "real_lemma_change_eq_subj"; | 
| 948 | ||
| 949 | (* FIXME: long! *) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 950 | Goal "(0r < S + %~W) ==> (W < S)"; | 
| 5078 | 951 | by (rtac ccontr 1); | 
| 952 | by (dtac (real_leI RS real_le_imp_less_or_eq) 1); | |
| 953 | by (auto_tac (claset(), | |
| 954 | simpset() addsimps [real_less_not_refl,real_add_minus])); | |
| 955 | by (EVERY1[dtac real_less_add_positive_left_Ex, etac exE, etac conjE]); | |
| 956 | by (asm_full_simp_tac (simpset() addsimps [real_add_zero_left]) 1); | |
| 957 | by (dtac real_lemma_change_eq_subj 1); | |
| 958 | by (auto_tac (claset(),simpset() addsimps [real_minus_minus])); | |
| 959 | by (dtac real_less_sum_gt_zero 1); | |
| 960 | by (asm_full_simp_tac (simpset() addsimps [real_minus_add_eq] @ real_add_ac) 1); | |
| 961 | by (EVERY1[rotate_tac 1, dtac (real_add_left_commute RS ssubst)]); | |
| 962 | by (auto_tac (claset() addEs [real_less_asym], | |
| 963 | simpset() addsimps [real_add_minus,real_add_zero_right])); | |
| 964 | qed "real_sum_gt_zero_less"; | |
| 965 | ||
| 966 | Goal "(0r < S + %~W) = (W < S)"; | |
| 967 | by (fast_tac (claset() addIs [real_less_sum_gt_zero, | |
| 968 | real_sum_gt_zero_less]) 1); | |
| 969 | qed "real_less_sum_gt_0_iff"; | |
| 970 | ||
| 971 | Goal "((x::real) < y) = (%~y < %~x)"; | |
| 972 | by (rtac (real_less_sum_gt_0_iff RS subst) 1); | |
| 973 | by (res_inst_tac [("W1","x")] (real_less_sum_gt_0_iff RS subst) 1);
 | |
| 974 | by (simp_tac (simpset() addsimps [real_add_commute]) 1); | |
| 975 | qed "real_less_swap_iff"; | |
| 976 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 977 | Goal "[| R + L = S; 0r < L |] ==> R < S"; | 
| 5078 | 978 | by (rtac (real_less_sum_gt_0_iff RS iffD1) 1); | 
| 979 | by (auto_tac (claset(),simpset() addsimps [ | |
| 980 | real_add_minus,real_add_zero_right] @ real_add_ac)); | |
| 981 | qed "real_lemma_add_positive_imp_less"; | |
| 982 | ||
| 983 | Goal "!!(R::real). ? T. 0r < T & R + T = S ==> R < S"; | |
| 984 | by (blast_tac (claset() addIs [real_lemma_add_positive_imp_less]) 1); | |
| 985 | qed "real_ex_add_positive_left_less"; | |
| 986 | ||
| 987 | (*** alternative definition for real_less ***) | |
| 988 | Goal "!!(R::real). (? T. 0r < T & R + T = S) = (R < S)"; | |
| 989 | by (fast_tac (claset() addSIs [real_less_add_positive_left_Ex, | |
| 990 | real_ex_add_positive_left_less]) 1); | |
| 991 | qed "real_less_iffdef"; | |
| 992 | ||
| 993 | Goal "(0r < x) = (%~x < x)"; | |
| 994 | by (Step_tac 1); | |
| 995 | by (rtac ccontr 2 THEN forward_tac | |
| 996 | [real_leI RS real_le_imp_less_or_eq] 2); | |
| 997 | by (Step_tac 2); | |
| 998 | by (dtac (real_minus_zero_less_iff RS iffD2) 2); | |
| 999 | by (fast_tac (claset() addDs [real_less_trans]) 2); | |
| 1000 | by (auto_tac (claset(),simpset() addsimps | |
| 1001 | [real_gt_zero_preal_Ex,real_preal_minus_less_self])); | |
| 1002 | qed "real_gt_zero_iff"; | |
| 1003 | ||
| 1004 | Goal "(x < 0r) = (x < %~x)"; | |
| 1005 | by (rtac (real_minus_zero_less_iff RS subst) 1); | |
| 1006 | by (stac real_gt_zero_iff 1); | |
| 1007 | by (Full_simp_tac 1); | |
| 1008 | qed "real_lt_zero_iff"; | |
| 1009 | ||
| 1010 | Goalw [real_le_def] "(0r <= x) = (%~x <= x)"; | |
| 1011 | by (auto_tac (claset(),simpset() addsimps [real_lt_zero_iff RS sym])); | |
| 1012 | qed "real_ge_zero_iff"; | |
| 1013 | ||
| 1014 | Goalw [real_le_def] "(x <= 0r) = (x <= %~x)"; | |
| 1015 | by (auto_tac (claset(),simpset() addsimps [real_gt_zero_iff RS sym])); | |
| 1016 | qed "real_le_zero_iff"; | |
| 1017 | ||
| 1018 | Goal "(%#m1 <= %#m2) = (m1 <= m2)"; | |
| 1019 | by (auto_tac (claset() addSIs [preal_leI], | |
| 1020 | simpset() addsimps [real_less_le_iff RS sym])); | |
| 1021 | by (dtac preal_le_less_trans 1 THEN assume_tac 1); | |
| 1022 | by (etac preal_less_irrefl 1); | |
| 1023 | qed "real_preal_le_iff"; | |
| 1024 | ||
| 1025 | Goal "!!(x::real). [| 0r < x; 0r < y |] ==> 0r < x * y"; | |
| 1026 | by (auto_tac (claset(),simpset() addsimps [real_gt_zero_preal_Ex])); | |
| 1027 | by (res_inst_tac [("x","y*ya")] exI 1);
 | |
| 1028 | by (full_simp_tac (simpset() addsimps [real_preal_mult]) 1); | |
| 1029 | qed "real_mult_order"; | |
| 1030 | ||
| 1031 | Goal "!!(x::real). [| x < 0r; y < 0r |] ==> 0r < x * y"; | |
| 1032 | by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1)); | |
| 1033 | by (dtac real_mult_order 1 THEN assume_tac 1); | |
| 1034 | by (Asm_full_simp_tac 1); | |
| 1035 | qed "real_mult_less_zero1"; | |
| 1036 | ||
| 1037 | Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x * y"; | |
| 1038 | by (REPEAT(dtac real_le_imp_less_or_eq 1)); | |
| 1039 | by (auto_tac (claset() addIs [real_mult_order, | |
| 1040 | real_less_imp_le],simpset() addsimps [real_le_refl])); | |
| 1041 | qed "real_le_mult_order"; | |
| 1042 | ||
| 1043 | Goal "!!(x::real). [| x <= 0r; y <= 0r |] ==> 0r <= x * y"; | |
| 1044 | by (rtac real_less_or_eq_imp_le 1); | |
| 1045 | by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1); | |
| 1046 | by Auto_tac; | |
| 1047 | by (dtac real_le_imp_less_or_eq 1); | |
| 1048 | by (auto_tac (claset() addDs [real_mult_less_zero1],simpset())); | |
| 1049 | qed "real_mult_le_zero1"; | |
| 1050 | ||
| 1051 | Goal "!!(x::real). [| 0r <= x; y < 0r |] ==> x * y <= 0r"; | |
| 1052 | by (rtac real_less_or_eq_imp_le 1); | |
| 1053 | by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1); | |
| 1054 | by Auto_tac; | |
| 1055 | by (dtac (real_minus_zero_less_iff RS iffD2) 1); | |
| 1056 | by (rtac (real_minus_zero_less_iff RS subst) 1); | |
| 1057 | by (blast_tac (claset() addDs [real_mult_order] | |
| 1058 | addIs [real_minus_mult_eq2 RS ssubst]) 1); | |
| 1059 | qed "real_mult_le_zero"; | |
| 1060 | ||
| 1061 | Goal "!!(x::real). [| 0r < x; y < 0r |] ==> x*y < 0r"; | |
| 1062 | by (dtac (real_minus_zero_less_iff RS iffD2) 1); | |
| 1063 | by (dtac real_mult_order 1 THEN assume_tac 1); | |
| 1064 | by (rtac (real_minus_zero_less_iff RS iffD1) 1); | |
| 1065 | by (asm_full_simp_tac (simpset() addsimps [real_minus_mult_eq2]) 1); | |
| 1066 | qed "real_mult_less_zero"; | |
| 1067 | ||
| 1068 | Goalw [real_one_def] "0r < 1r"; | |
| 1069 | by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2], | |
| 1070 | simpset() addsimps [real_preal_def])); | |
| 1071 | qed "real_zero_less_one"; | |
| 1072 | ||
| 1073 | (*** Completeness of reals ***) | |
| 1074 | (** use supremum property of preal and theorems about real_preal **) | |
| 1075 | (*** a few lemmas ***) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1076 | Goal "! x:P. 0r < x ==> ((? x:P. y < x) = (? X. %#X : P & y < %#X))"; | 
| 5078 | 1077 | by (blast_tac (claset() addSDs [bspec,real_gt_zero_preal_Ex RS iffD1]) 1); | 
| 1078 | qed "real_sup_lemma1"; | |
| 1079 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1080 | Goal "[| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \ | 
| 5078 | 1081 | \         ==> (? X. X: {w. %#w : P}) & (? Y. !X: {w. %#w : P}. X < Y)";
 | 
| 1082 | by (rtac conjI 1); | |
| 1083 | by (blast_tac (claset() addDs [bspec,real_gt_zero_preal_Ex RS iffD1]) 1); | |
| 1084 | by Auto_tac; | |
| 1085 | by (dtac bspec 1 THEN assume_tac 1); | |
| 1086 | by (forward_tac [bspec] 1 THEN assume_tac 1); | |
| 1087 | by (dtac real_less_trans 1 THEN assume_tac 1); | |
| 1088 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1 THEN etac exE 1); | |
| 1089 | by (res_inst_tac [("x","ya")] exI 1);
 | |
| 1090 | by Auto_tac; | |
| 1091 | by (dres_inst_tac [("x","%#X")] bspec 1 THEN assume_tac 1);
 | |
| 1092 | by (etac real_preal_lessD 1); | |
| 1093 | qed "real_sup_lemma2"; | |
| 1094 | ||
| 1095 | (*------------------------------------------------------------- | |
| 1096 | Completeness of Positive Reals | |
| 1097 | -------------------------------------------------------------*) | |
| 1098 | ||
| 1099 | (* Supremum property for the set of positive reals *) | |
| 1100 | (* FIXME: long proof - can be improved - need only have one case split *) | |
| 1101 | (* will do for now *) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1102 | Goal "[| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \ | 
| 5078 | 1103 | \ ==> (? S. !y. (? x: P. y < x) = (y < S))"; | 
| 1104 | by (res_inst_tac [("x","%#psup({w. %#w : P})")] exI 1);
 | |
| 1105 | by Auto_tac; | |
| 1106 | by (forward_tac [real_sup_lemma2] 1 THEN Auto_tac); | |
| 1107 | by (case_tac "0r < ya" 1); | |
| 1108 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); | |
| 1109 | by (dtac real_less_all_real2 2); | |
| 1110 | by Auto_tac; | |
| 1111 | by (rtac (preal_complete RS spec RS iffD1) 1); | |
| 1112 | by Auto_tac; | |
| 1113 | by (forward_tac [real_gt_preal_preal_Ex] 1); | |
| 1114 | by Auto_tac; | |
| 1115 | (* second part *) | |
| 1116 | by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1); | |
| 1117 | by (case_tac "0r < ya" 1); | |
| 1118 | by (auto_tac (claset() addSDs [real_less_all_real2, | |
| 1119 | real_gt_zero_preal_Ex RS iffD1],simpset())); | |
| 1120 | by (forward_tac [real_sup_lemma2] 2 THEN Auto_tac); | |
| 1121 | by (forward_tac [real_sup_lemma2] 1 THEN Auto_tac); | |
| 1122 | by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1); | |
| 1123 | by (Fast_tac 3); | |
| 1124 | by (Fast_tac 1); | |
| 1125 | by (Fast_tac 1); | |
| 1126 | by (Blast_tac 1); | |
| 1127 | qed "posreal_complete"; | |
| 1128 | ||
| 1129 | (*------------------------------------------------------------------ | |
| 1130 | ||
| 1131 | ------------------------------------------------------------------*) | |
| 1132 | ||
| 1133 | Goal "!!(A::real). A < B ==> A + C < B + C"; | |
| 1134 | by (dtac (real_less_iffdef RS iffD2) 1); | |
| 1135 | by (rtac (real_less_iffdef RS iffD1) 1); | |
| 1136 | by (REPEAT(Step_tac 1)); | |
| 1137 | by (full_simp_tac (simpset() addsimps real_add_ac) 1); | |
| 1138 | qed "real_add_less_mono1"; | |
| 1139 | ||
| 1140 | Goal "!!(A::real). A < B ==> C + A < C + B"; | |
| 1141 | by (auto_tac (claset() addIs [real_add_less_mono1], | |
| 1142 | simpset() addsimps [real_add_commute])); | |
| 1143 | qed "real_add_less_mono2"; | |
| 1144 | ||
| 1145 | Goal "!!(A::real). A + C < B + C ==> A < B"; | |
| 1146 | by (dres_inst_tac [("C","%~C")] real_add_less_mono1 1);
 | |
| 1147 | by (asm_full_simp_tac (simpset() addsimps [real_add_assoc, | |
| 1148 | real_add_minus,real_add_zero_right]) 1); | |
| 1149 | qed "real_less_add_right_cancel"; | |
| 1150 | ||
| 1151 | Goal "!!(A::real). C + A < C + B ==> A < B"; | |
| 1152 | by (dres_inst_tac [("C","%~C")] real_add_less_mono2 1);
 | |
| 1153 | by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym, | |
| 1154 | real_add_minus_left,real_add_zero_left]) 1); | |
| 1155 | qed "real_less_add_left_cancel"; | |
| 1156 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1157 | Goal "[| 0r < x; 0r < y |] ==> 0r < x + y"; | 
| 5078 | 1158 | by (REPEAT(dtac (real_gt_zero_preal_Ex RS iffD1) 1)); | 
| 1159 | by (rtac (real_gt_zero_preal_Ex RS iffD2) 1); | |
| 1160 | by (Step_tac 1); | |
| 1161 | by (res_inst_tac [("x","y + ya")] exI 1);
 | |
| 1162 | by (full_simp_tac (simpset() addsimps [real_preal_add]) 1); | |
| 1163 | qed "real_add_order"; | |
| 1164 | ||
| 1165 | Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x + y"; | |
| 1166 | by (REPEAT(dtac real_le_imp_less_or_eq 1)); | |
| 1167 | by (auto_tac (claset() addIs [real_add_order, | |
| 1168 | real_less_imp_le],simpset() addsimps [real_add_zero_left, | |
| 1169 | real_add_zero_right,real_le_refl])); | |
| 1170 | qed "real_le_add_order"; | |
| 1171 | ||
| 1172 | Goal | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 1173 | "[| R1 < S1; R2 < S2 |] ==> R1 + R2 < S1 + (S2::real)"; | 
| 5078 | 1174 | by (dtac (real_less_iffdef RS iffD2) 1); | 
| 1175 | by (dtac (real_less_iffdef RS iffD2) 1); | |
| 1176 | by (rtac (real_less_iffdef RS iffD1) 1); | |
| 1177 | by Auto_tac; | |
| 1178 | by (res_inst_tac [("x","T + Ta")] exI 1);
 | |
| 1179 | by (auto_tac (claset(),simpset() addsimps [real_add_order] @ real_add_ac)); | |
| 1180 | qed "real_add_less_mono"; | |
| 1181 | ||
| 1182 | Goal "!!(x::real). [| 0r <= x; 0r <= y |] ==> 0r <= x + y"; | |
| 1183 | by (REPEAT(dtac real_le_imp_less_or_eq 1)); | |
| 1184 | by (auto_tac (claset() addIs [real_add_order, | |
| 1185 | real_less_imp_le],simpset() addsimps [real_add_zero_left, | |
| 1186 | real_add_zero_right,real_le_refl])); | |
| 1187 | qed "real_le_add_order"; | |
| 1188 | ||
| 1189 | Goal "!!(q1::real). q1 <= q2 ==> x + q1 <= x + q2"; | |
| 1190 | by (dtac real_le_imp_less_or_eq 1); | |
| 1191 | by (Step_tac 1); | |
| 1192 | by (auto_tac (claset() addSIs [real_le_refl, | |
| 1193 | real_less_imp_le,real_add_less_mono1], | |
| 1194 | simpset() addsimps [real_add_commute])); | |
| 1195 | qed "real_add_left_le_mono1"; | |
| 1196 | ||
| 1197 | Goal "!!(q1::real). q1 <= q2 ==> q1 + x <= q2 + x"; | |
| 1198 | by (auto_tac (claset() addDs [real_add_left_le_mono1], | |
| 1199 | simpset() addsimps [real_add_commute])); | |
| 1200 | qed "real_add_le_mono1"; | |
| 1201 | ||
| 1202 | Goal "!!k l::real. [|i<=j; k<=l |] ==> i + k <= j + l"; | |
| 1203 | by (etac (real_add_le_mono1 RS real_le_trans) 1); | |
| 1204 | by (simp_tac (simpset() addsimps [real_add_commute]) 1); | |
| 1205 | (*j moves to the end because it is free while k, l are bound*) | |
| 1206 | by (etac real_add_le_mono1 1); | |
| 1207 | qed "real_add_le_mono"; | |
| 1208 | ||
| 1209 | Goal "EX (x::real). x < y"; | |
| 1210 | by (rtac (real_add_zero_right RS subst) 1); | |
| 1211 | by (res_inst_tac [("x","y + %~1r")] exI 1);
 | |
| 1212 | by (auto_tac (claset() addSIs [real_add_less_mono2], | |
| 1213 | simpset() addsimps [real_minus_zero_less_iff2, | |
| 1214 | real_zero_less_one])); | |
| 1215 | qed "real_less_Ex"; | |
| 1216 | (*--------------------------------------------------------------------------------- | |
| 1217 | An embedding of the naturals in the reals | |
| 1218 | ---------------------------------------------------------------------------------*) | |
| 1219 | ||
| 1220 | Goalw [real_nat_def] "%%#0 = 1r"; | |
| 1221 | by (full_simp_tac (simpset() addsimps [pnat_one_iff RS sym,real_preal_def]) 1); | |
| 1222 | by (fold_tac [real_one_def]); | |
| 1223 | by (rtac refl 1); | |
| 1224 | qed "real_nat_one"; | |
| 1225 | ||
| 1226 | Goalw [real_nat_def] "%%#1 = 1r + 1r"; | |
| 1227 | by (full_simp_tac (simpset() addsimps [real_preal_def,real_one_def, | |
| 1228 | pnat_two_eq,real_add,prat_pnat_add RS sym,preal_prat_add RS sym | |
| 1229 | ] @ pnat_add_ac) 1); | |
| 1230 | qed "real_nat_two"; | |
| 1231 | ||
| 1232 | Goalw [real_nat_def] | |
| 1233 | "%%#n1 + %%#n2 = %%#(n1 + n2) + 1r"; | |
| 1234 | by (full_simp_tac (simpset() addsimps [real_nat_one RS sym, | |
| 1235 | real_nat_def,real_preal_add RS sym,preal_prat_add RS sym, | |
| 1236 | prat_pnat_add RS sym,pnat_nat_add]) 1); | |
| 1237 | qed "real_nat_add"; | |
| 1238 | ||
| 1239 | Goal "%%#(n + 1) = %%#n + 1r"; | |
| 1240 | by (res_inst_tac [("x1","1r")] (real_add_right_cancel RS iffD1) 1);
 | |
| 1241 | by (rtac (real_nat_add RS subst) 1); | |
| 1242 | by (full_simp_tac (simpset() addsimps [real_nat_two,real_add_assoc]) 1); | |
| 1243 | qed "real_nat_add_one"; | |
| 1244 | ||
| 1245 | Goal "Suc n = n + 1"; | |
| 1246 | by Auto_tac; | |
| 1247 | qed "lemma"; | |
| 1248 | ||
| 1249 | Goal "%%#Suc n = %%#n + 1r"; | |
| 1250 | by (stac lemma 1); | |
| 1251 | by (rtac real_nat_add_one 1); | |
| 1252 | qed "real_nat_Suc"; | |
| 1253 | ||
| 1254 | Goal "inj(real_nat)"; | |
| 1255 | by (rtac injI 1); | |
| 1256 | by (rewtac real_nat_def); | |
| 1257 | by (dtac (inj_real_preal RS injD) 1); | |
| 1258 | by (dtac (inj_preal_prat RS injD) 1); | |
| 1259 | by (dtac (inj_prat_pnat RS injD) 1); | |
| 1260 | by (etac (inj_pnat_nat RS injD) 1); | |
| 1261 | qed "inj_real_nat"; | |
| 1262 | ||
| 1263 | Goalw [real_nat_def] "0r < %%#n"; | |
| 1264 | by (rtac (real_gt_zero_preal_Ex RS iffD2) 1); | |
| 1265 | by (Blast_tac 1); | |
| 1266 | qed "real_nat_less_zero"; | |
| 1267 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1268 | Goal "1r <= %%#n"; | 
| 5078 | 1269 | by (simp_tac (simpset() addsimps [real_nat_one RS sym]) 1); | 
| 5184 | 1270 | by (induct_tac "n" 1); | 
| 5078 | 1271 | by (auto_tac (claset(),simpset () | 
| 1272 | addsimps [real_nat_Suc,real_le_refl,real_nat_one])); | |
| 1273 | by (res_inst_tac [("t","1r")] (real_add_zero_left RS subst) 1);
 | |
| 1274 | by (rtac real_add_le_mono 1); | |
| 1275 | by (auto_tac (claset(),simpset () | |
| 1276 | addsimps [real_le_refl,real_nat_less_zero, | |
| 1277 | real_less_imp_le,real_add_zero_left])); | |
| 1278 | qed "real_nat_less_one"; | |
| 1279 | ||
| 1280 | Goal "rinv(%%#n) ~= 0r"; | |
| 1281 | by (rtac ((real_nat_less_zero RS | |
| 1282 | real_not_refl2 RS not_sym) RS rinv_not_zero) 1); | |
| 1283 | qed "real_nat_rinv_not_zero"; | |
| 1284 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1285 | Goal "rinv(%%#x) = rinv(%%#y) ==> x = y"; | 
| 5078 | 1286 | by (rtac (inj_real_nat RS injD) 1); | 
| 1287 | by (res_inst_tac [("n2","x")] 
 | |
| 1288 | (real_nat_rinv_not_zero RS real_mult_left_cancel RS iffD1) 1); | |
| 1289 | by (full_simp_tac (simpset() addsimps [(real_nat_less_zero RS | |
| 1290 | real_not_refl2 RS not_sym) RS real_mult_inv_left]) 1); | |
| 1291 | by (asm_full_simp_tac (simpset() addsimps [(real_nat_less_zero RS | |
| 1292 | real_not_refl2 RS not_sym)]) 1); | |
| 1293 | qed "real_nat_rinv_inj"; | |
| 1294 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1295 | Goal "0r < x ==> 0r < rinv x"; | 
| 5078 | 1296 | by (EVERY1[rtac ccontr, dtac real_leI]); | 
| 1297 | by (forward_tac [real_minus_zero_less_iff2 RS iffD2] 1); | |
| 1298 | by (forward_tac [real_not_refl2 RS not_sym] 1); | |
| 1299 | by (dtac (real_not_refl2 RS not_sym RS rinv_not_zero) 1); | |
| 1300 | by (EVERY1[dtac real_le_imp_less_or_eq, Step_tac]); | |
| 1301 | by (dtac real_mult_less_zero1 1 THEN assume_tac 1); | |
| 1302 | by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym], | |
| 1303 | simpset() addsimps [real_minus_mult_eq1 RS sym])); | |
| 1304 | qed "real_rinv_gt_zero"; | |
| 1305 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1306 | Goal "x < 0r ==> rinv x < 0r"; | 
| 5078 | 1307 | by (forward_tac [real_not_refl2] 1); | 
| 1308 | by (dtac (real_minus_zero_less_iff RS iffD2) 1); | |
| 1309 | by (rtac (real_minus_zero_less_iff RS iffD1) 1); | |
| 1310 | by (dtac (real_minus_rinv RS sym) 1); | |
| 1311 | by (auto_tac (claset() addIs [real_rinv_gt_zero], | |
| 1312 | simpset())); | |
| 1313 | qed "real_rinv_less_zero"; | |
| 1314 | ||
| 1315 | Goal "x+x=x*(1r+1r)"; | |
| 1316 | by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1); | |
| 1317 | qed "real_add_self"; | |
| 1318 | ||
| 1319 | Goal "x < x + 1r"; | |
| 1320 | by (rtac (real_less_sum_gt_0_iff RS iffD1) 1); | |
| 1321 | by (full_simp_tac (simpset() addsimps [real_zero_less_one, | |
| 1322 | real_add_assoc,real_add_minus,real_add_zero_right, | |
| 1323 | real_add_left_commute]) 1); | |
| 1324 | qed "real_self_less_add_one"; | |
| 1325 | ||
| 1326 | Goal "1r < 1r + 1r"; | |
| 1327 | by (rtac real_self_less_add_one 1); | |
| 1328 | qed "real_one_less_two"; | |
| 1329 | ||
| 1330 | Goal "0r < 1r + 1r"; | |
| 1331 | by (rtac ([real_zero_less_one, | |
| 1332 | real_one_less_two] MRS real_less_trans) 1); | |
| 1333 | qed "real_zero_less_two"; | |
| 1334 | ||
| 1335 | Goal "1r + 1r ~= 0r"; | |
| 1336 | by (rtac (real_zero_less_two RS real_not_refl2 RS not_sym) 1); | |
| 1337 | qed "real_two_not_zero"; | |
| 1338 | ||
| 1339 | Addsimps [real_two_not_zero]; | |
| 1340 | ||
| 1341 | Goal "x*rinv(1r + 1r) + x*rinv(1r + 1r) = x"; | |
| 1342 | by (stac real_add_self 1); | |
| 1343 | by (full_simp_tac (simpset() addsimps [real_mult_assoc]) 1); | |
| 1344 | qed "real_sum_of_halves"; | |
| 1345 | ||
| 1346 | Goal "!!(x::real). [| 0r<z; x<y |] ==> x*z<y*z"; | |
| 1347 | by (rotate_tac 1 1); | |
| 1348 | by (dtac real_less_sum_gt_zero 1); | |
| 1349 | by (rtac real_sum_gt_zero_less 1); | |
| 1350 | by (dtac real_mult_order 1 THEN assume_tac 1); | |
| 1351 | by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2, | |
| 1352 | real_minus_mult_eq2 RS sym, real_mult_commute ]) 1); | |
| 1353 | qed "real_mult_less_mono1"; | |
| 1354 | ||
| 1355 | Goal "!!(y::real). [| 0r<z; x<y |] ==> z*x<z*y"; | |
| 1356 | by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_less_mono1]) 1); | |
| 1357 | qed "real_mult_less_mono2"; | |
| 1358 | ||
| 1359 | Goal "!!(x::real). [| 0r<z; x*z<y*z |] ==> x<y"; | |
| 1360 | by (forw_inst_tac [("x","x*z")] (real_rinv_gt_zero 
 | |
| 1361 | RS real_mult_less_mono1) 1); | |
| 1362 | by (auto_tac (claset(),simpset() addsimps | |
| 1363 | [real_mult_assoc,real_not_refl2 RS not_sym])); | |
| 1364 | qed "real_mult_less_cancel1"; | |
| 1365 | ||
| 1366 | Goal "!!(x::real). [| 0r<z; z*x<z*y |] ==> x<y"; | |
| 1367 | by (etac real_mult_less_cancel1 1); | |
| 1368 | by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1); | |
| 1369 | qed "real_mult_less_cancel2"; | |
| 1370 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1371 | Goal "0r < z ==> (x*z < y*z) = (x < y)"; | 
| 5078 | 1372 | by (blast_tac (claset() addIs [real_mult_less_mono1, | 
| 1373 | real_mult_less_cancel1]) 1); | |
| 1374 | qed "real_mult_less_iff1"; | |
| 1375 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1376 | Goal "0r < z ==> (z*x < z*y) = (x < y)"; | 
| 5078 | 1377 | by (blast_tac (claset() addIs [real_mult_less_mono2, | 
| 1378 | real_mult_less_cancel2]) 1); | |
| 1379 | qed "real_mult_less_iff2"; | |
| 1380 | ||
| 1381 | Addsimps [real_mult_less_iff1,real_mult_less_iff2]; | |
| 1382 | ||
| 1383 | Goal "!!(x::real). [| 0r<=z; x<y |] ==> x*z<=y*z"; | |
| 1384 | by (EVERY1 [rtac real_less_or_eq_imp_le, dtac real_le_imp_less_or_eq]); | |
| 1385 | by (auto_tac (claset() addIs [real_mult_less_mono1],simpset())); | |
| 1386 | qed "real_mult_le_less_mono1"; | |
| 1387 | ||
| 1388 | Goal "!!(x::real). [| 0r<=z; x<y |] ==> z*x<=z*y"; | |
| 1389 | by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_le_less_mono1]) 1); | |
| 1390 | qed "real_mult_le_less_mono2"; | |
| 1391 | ||
| 1392 | Goal "!!x y (z::real). [| 0r<=z; x<=y |] ==> z*x<=z*y"; | |
| 1393 | by (dres_inst_tac [("x","x")] real_le_imp_less_or_eq 1);
 | |
| 1394 | by (auto_tac (claset() addIs [real_mult_le_less_mono2,real_le_refl],simpset())); | |
| 1395 | qed "real_mult_le_le_mono1"; | |
| 1396 | ||
| 1397 | Goal "!!(x::real). x < y ==> x < (x + y)*rinv(1r + 1r)"; | |
| 1398 | by (dres_inst_tac [("C","x")] real_add_less_mono2 1);
 | |
| 1399 | by (dtac (real_add_self RS subst) 1); | |
| 1400 | by (dtac (real_zero_less_two RS real_rinv_gt_zero RS | |
| 1401 | real_mult_less_mono1) 1); | |
| 1402 | by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1); | |
| 1403 | qed "real_less_half_sum"; | |
| 1404 | ||
| 1405 | Goal "!!(x::real). x < y ==> (x + y)*rinv(1r + 1r) < y"; | |
| 1406 | by (dres_inst_tac [("C","y")] real_add_less_mono1 1);
 | |
| 1407 | by (dtac (real_add_self RS subst) 1); | |
| 1408 | by (dtac (real_zero_less_two RS real_rinv_gt_zero RS | |
| 1409 | real_mult_less_mono1) 1); | |
| 1410 | by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1); | |
| 1411 | qed "real_gt_half_sum"; | |
| 1412 | ||
| 1413 | Goal "!!(x::real). x < y ==> EX r. x < r & r < y"; | |
| 1414 | by (blast_tac (claset() addSIs [real_less_half_sum,real_gt_half_sum]) 1); | |
| 1415 | qed "real_dense"; | |
| 1416 | ||
| 1417 | Goal "(EX n. rinv(%%#n) < r) = (EX n. 1r < r * %%#n)"; | |
| 1418 | by (Step_tac 1); | |
| 1419 | by (dres_inst_tac [("n1","n")] (real_nat_less_zero 
 | |
| 1420 | RS real_mult_less_mono1) 1); | |
| 1421 | by (dres_inst_tac [("n2","n")] (real_nat_less_zero RS 
 | |
| 1422 | real_rinv_gt_zero RS real_mult_less_mono1) 2); | |
| 1423 | by (auto_tac (claset(),simpset() addsimps [(real_nat_less_zero RS | |
| 1424 | real_not_refl2 RS not_sym),real_mult_assoc])); | |
| 1425 | qed "real_nat_rinv_Ex_iff"; | |
| 1426 | ||
| 1427 | Goalw [real_nat_def] "(%%#n < %%#m) = (n < m)"; | |
| 1428 | by Auto_tac; | |
| 1429 | qed "real_nat_less_iff"; | |
| 1430 | ||
| 1431 | Addsimps [real_nat_less_iff]; | |
| 1432 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1433 | Goal "0r < u ==> (u < rinv (%%#n)) = (%%#n < rinv(u))"; | 
| 5078 | 1434 | by (Step_tac 1); | 
| 1435 | by (res_inst_tac [("n2","n")] (real_nat_less_zero RS 
 | |
| 1436 | real_rinv_gt_zero RS real_mult_less_cancel1) 1); | |
| 1437 | by (res_inst_tac [("x1","u")] ( real_rinv_gt_zero
 | |
| 1438 | RS real_mult_less_cancel1) 2); | |
| 1439 | by (auto_tac (claset(),simpset() addsimps [real_nat_less_zero, | |
| 1440 | real_not_refl2 RS not_sym])); | |
| 1441 | by (res_inst_tac [("z","u")] real_mult_less_cancel2 1);
 | |
| 1442 | by (res_inst_tac [("n1","n")] (real_nat_less_zero RS 
 | |
| 1443 | real_mult_less_cancel2) 3); | |
| 1444 | by (auto_tac (claset(),simpset() addsimps [real_nat_less_zero, | |
| 1445 | real_not_refl2 RS not_sym,real_mult_assoc RS sym])); | |
| 1446 | qed "real_nat_less_rinv_iff"; | |
| 1447 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 1448 | Goal "0r < u ==> (u = rinv(%%#n)) = (%%#n = rinv u)"; | 
| 5078 | 1449 | by (auto_tac (claset(),simpset() addsimps [real_rinv_rinv, | 
| 1450 | real_nat_less_zero,real_not_refl2 RS not_sym])); | |
| 1451 | qed "real_nat_rinv_eq_iff"; | |
| 1452 | ||
| 1453 | (* | |
| 1454 | (*------------------------------------------------------------------ | |
| 1455 | lemmas about upper bounds and least upper bound | |
| 1456 | ------------------------------------------------------------------*) | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 1457 | Goalw [real_ub_def] "[| real_ub u S; x : S |] ==> x <= u"; | 
| 5078 | 1458 | by Auto_tac; | 
| 1459 | qed "real_ubD"; | |
| 1460 | ||
| 1461 | *) |