src/HOL/Real/RealDef.ML
author paulson
Tue, 23 Nov 1999 11:19:39 +0100
changeset 8027 8a27d0579e37
parent 7499 23e090051cb8
child 9043 ca761fe227d8
permissions -rw-r--r--
distributive laws for * over -
     1 (*  Title       : Real/RealDef.ML
     2     ID          : $Id$
     3     Author      : Jacques D. Fleuriot
     4     Copyright   : 1998  University of Cambridge
     5     Description : The reals
     6 *)
     7 
     8 (*** Proving that realrel is an equivalence relation ***)
     9 
    10 Goal "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] \
    11 \            ==> x1 + y3 = x3 + y1";        
    12 by (res_inst_tac [("C","y2")] preal_add_right_cancel 1);
    13 by (rotate_tac 1 1 THEN dtac sym 1);
    14 by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
    15 by (rtac (preal_add_left_commute RS subst) 1);
    16 by (res_inst_tac [("x1","x1")] (preal_add_assoc RS subst) 1);
    17 by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
    18 qed "preal_trans_lemma";
    19 
    20 (** Natural deduction for realrel **)
    21 
    22 Goalw [realrel_def]
    23     "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)";
    24 by (Blast_tac 1);
    25 qed "realrel_iff";
    26 
    27 Goalw [realrel_def]
    28     "[| x1 + y2 = x2 + y1 |] ==> ((x1,y1),(x2,y2)): realrel";
    29 by (Blast_tac  1);
    30 qed "realrelI";
    31 
    32 Goalw [realrel_def]
    33   "p: realrel --> (EX x1 y1 x2 y2. \
    34 \                  p = ((x1,y1),(x2,y2)) & x1 + y2 = x2 + y1)";
    35 by (Blast_tac 1);
    36 qed "realrelE_lemma";
    37 
    38 val [major,minor] = goal thy
    39   "[| p: realrel;  \
    40 \     !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2));  x1+y2 = x2+y1 \
    41 \                    |] ==> Q |] ==> Q";
    42 by (cut_facts_tac [major RS (realrelE_lemma RS mp)] 1);
    43 by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
    44 qed "realrelE";
    45 
    46 AddSIs [realrelI];
    47 AddSEs [realrelE];
    48 
    49 Goal "(x,x): realrel";
    50 by (stac surjective_pairing 1 THEN rtac (refl RS realrelI) 1);
    51 qed "realrel_refl";
    52 
    53 Goalw [equiv_def, refl_def, sym_def, trans_def]
    54     "equiv {x::(preal*preal).True} realrel";
    55 by (fast_tac (claset() addSIs [realrel_refl] 
    56                       addSEs [sym,preal_trans_lemma]) 1);
    57 qed "equiv_realrel";
    58 
    59 val equiv_realrel_iff =
    60     [TrueI, TrueI] MRS 
    61     ([CollectI, CollectI] MRS 
    62     (equiv_realrel RS eq_equiv_class_iff));
    63 
    64 Goalw  [real_def,realrel_def,quotient_def] "realrel^^{(x,y)}:real";
    65 by (Blast_tac 1);
    66 qed "realrel_in_real";
    67 
    68 Goal "inj_on Abs_real real";
    69 by (rtac inj_on_inverseI 1);
    70 by (etac Abs_real_inverse 1);
    71 qed "inj_on_Abs_real";
    72 
    73 Addsimps [equiv_realrel_iff,inj_on_Abs_real RS inj_on_iff,
    74           realrel_iff, realrel_in_real, Abs_real_inverse];
    75 
    76 Addsimps [equiv_realrel RS eq_equiv_class_iff];
    77 val eq_realrelD = equiv_realrel RSN (2,eq_equiv_class);
    78 
    79 Goal "inj(Rep_real)";
    80 by (rtac inj_inverseI 1);
    81 by (rtac Rep_real_inverse 1);
    82 qed "inj_Rep_real";
    83 
    84 (** real_of_preal: the injection from preal to real **)
    85 Goal "inj(real_of_preal)";
    86 by (rtac injI 1);
    87 by (rewtac real_of_preal_def);
    88 by (dtac (inj_on_Abs_real RS inj_onD) 1);
    89 by (REPEAT (rtac realrel_in_real 1));
    90 by (dtac eq_equiv_class 1);
    91 by (rtac equiv_realrel 1);
    92 by (Blast_tac 1);
    93 by Safe_tac;
    94 by (Asm_full_simp_tac 1);
    95 qed "inj_real_of_preal";
    96 
    97 val [prem] = goal thy
    98     "(!!x y. z = Abs_real(realrel^^{(x,y)}) ==> P) ==> P";
    99 by (res_inst_tac [("x1","z")] 
   100     (rewrite_rule [real_def] Rep_real RS quotientE) 1);
   101 by (dres_inst_tac [("f","Abs_real")] arg_cong 1);
   102 by (res_inst_tac [("p","x")] PairE 1);
   103 by (rtac prem 1);
   104 by (asm_full_simp_tac (simpset() addsimps [Rep_real_inverse]) 1);
   105 qed "eq_Abs_real";
   106 
   107 (**** real_minus: additive inverse on real ****)
   108 
   109 Goalw [congruent_def]
   110   "congruent realrel (%p. split (%x y. realrel^^{(y,x)}) p)";
   111 by Safe_tac;
   112 by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1);
   113 qed "real_minus_congruent";
   114 
   115 (*Resolve th against the corresponding facts for real_minus*)
   116 val real_minus_ize = RSLIST [equiv_realrel, real_minus_congruent];
   117 
   118 Goalw [real_minus_def]
   119       "- (Abs_real(realrel^^{(x,y)})) = Abs_real(realrel ^^ {(y,x)})";
   120 by (res_inst_tac [("f","Abs_real")] arg_cong 1);
   121 by (simp_tac (simpset() addsimps 
   122    [realrel_in_real RS Abs_real_inverse,real_minus_ize UN_equiv_class]) 1);
   123 qed "real_minus";
   124 
   125 Goal "- (- z) = (z::real)";
   126 by (res_inst_tac [("z","z")] eq_Abs_real 1);
   127 by (asm_simp_tac (simpset() addsimps [real_minus]) 1);
   128 qed "real_minus_minus";
   129 
   130 Addsimps [real_minus_minus];
   131 
   132 Goal "inj(%r::real. -r)";
   133 by (rtac injI 1);
   134 by (dres_inst_tac [("f","uminus")] arg_cong 1);
   135 by (asm_full_simp_tac (simpset() addsimps [real_minus_minus]) 1);
   136 qed "inj_real_minus";
   137 
   138 Goalw [real_zero_def] "-0r = 0r";
   139 by (simp_tac (simpset() addsimps [real_minus]) 1);
   140 qed "real_minus_zero";
   141 
   142 Addsimps [real_minus_zero];
   143 
   144 Goal "(-x = 0r) = (x = 0r)"; 
   145 by (res_inst_tac [("z","x")] eq_Abs_real 1);
   146 by (auto_tac (claset(),
   147 	      simpset() addsimps [real_zero_def, real_minus] @ preal_add_ac));
   148 qed "real_minus_zero_iff";
   149 
   150 Addsimps [real_minus_zero_iff];
   151 
   152 Goal "(-x ~= 0r) = (x ~= 0r)"; 
   153 by Auto_tac;
   154 qed "real_minus_not_zero_iff";
   155 
   156 (*** Congruence property for addition ***)
   157 Goalw [congruent2_def]
   158     "congruent2 realrel (%p1 p2.                  \
   159 \         split (%x1 y1. split (%x2 y2. realrel^^{(x1+x2, y1+y2)}) p2) p1)";
   160 by Safe_tac;
   161 by (asm_simp_tac (simpset() addsimps [preal_add_assoc]) 1);
   162 by (res_inst_tac [("z1.1","x1a")] (preal_add_left_commute RS ssubst) 1);
   163 by (asm_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
   164 by (asm_simp_tac (simpset() addsimps preal_add_ac) 1);
   165 qed "real_add_congruent2";
   166 
   167 (*Resolve th against the corresponding facts for real_add*)
   168 val real_add_ize = RSLIST [equiv_realrel, real_add_congruent2];
   169 
   170 Goalw [real_add_def]
   171   "Abs_real(realrel^^{(x1,y1)}) + Abs_real(realrel^^{(x2,y2)}) = \
   172 \  Abs_real(realrel^^{(x1+x2, y1+y2)})";
   173 by (asm_simp_tac (simpset() addsimps [real_add_ize UN_equiv_class2]) 1);
   174 qed "real_add";
   175 
   176 Goal "(z::real) + w = w + z";
   177 by (res_inst_tac [("z","z")] eq_Abs_real 1);
   178 by (res_inst_tac [("z","w")] eq_Abs_real 1);
   179 by (asm_simp_tac (simpset() addsimps preal_add_ac @ [real_add]) 1);
   180 qed "real_add_commute";
   181 
   182 Goal "((z1::real) + z2) + z3 = z1 + (z2 + z3)";
   183 by (res_inst_tac [("z","z1")] eq_Abs_real 1);
   184 by (res_inst_tac [("z","z2")] eq_Abs_real 1);
   185 by (res_inst_tac [("z","z3")] eq_Abs_real 1);
   186 by (asm_simp_tac (simpset() addsimps [real_add, preal_add_assoc]) 1);
   187 qed "real_add_assoc";
   188 
   189 (*For AC rewriting*)
   190 Goal "(x::real)+(y+z)=y+(x+z)";
   191 by (rtac (real_add_commute RS trans) 1);
   192 by (rtac (real_add_assoc RS trans) 1);
   193 by (rtac (real_add_commute RS arg_cong) 1);
   194 qed "real_add_left_commute";
   195 
   196 (* real addition is an AC operator *)
   197 bind_thms ("real_add_ac", [real_add_assoc,real_add_commute,real_add_left_commute]);
   198 
   199 Goalw [real_of_preal_def,real_zero_def] "0r + z = z";
   200 by (res_inst_tac [("z","z")] eq_Abs_real 1);
   201 by (asm_full_simp_tac (simpset() addsimps [real_add] @ preal_add_ac) 1);
   202 qed "real_add_zero_left";
   203 Addsimps [real_add_zero_left];
   204 
   205 Goal "z + 0r = z";
   206 by (simp_tac (simpset() addsimps [real_add_commute]) 1);
   207 qed "real_add_zero_right";
   208 Addsimps [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 Addsimps [real_add_minus];
   216 
   217 Goal "(-z) + z = 0r";
   218 by (simp_tac (simpset() addsimps [real_add_commute]) 1);
   219 qed "real_add_minus_left";
   220 Addsimps [real_add_minus_left];
   221 
   222 
   223 Goal "z + ((- z) + w) = (w::real)";
   224 by (simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
   225 qed "real_add_minus_cancel";
   226 
   227 Goal "(-z) + (z + w) = (w::real)";
   228 by (simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
   229 qed "real_minus_add_cancel";
   230 
   231 Addsimps [real_add_minus_cancel, real_minus_add_cancel];
   232 
   233 Goal "? y. (x::real) + y = 0r";
   234 by (blast_tac (claset() addIs [real_add_minus]) 1);
   235 qed "real_minus_ex";
   236 
   237 Goal "?! y. (x::real) + y = 0r";
   238 by (auto_tac (claset() addIs [real_add_minus],simpset()));
   239 by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
   240 by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
   241 by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
   242 qed "real_minus_ex1";
   243 
   244 Goal "?! y. y + (x::real) = 0r";
   245 by (auto_tac (claset() addIs [real_add_minus_left],simpset()));
   246 by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
   247 by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
   248 by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
   249 qed "real_minus_left_ex1";
   250 
   251 Goal "x + y = 0r ==> x = -y";
   252 by (cut_inst_tac [("z","y")] real_add_minus_left 1);
   253 by (res_inst_tac [("x1","y")] (real_minus_left_ex1 RS ex1E) 1);
   254 by (Blast_tac 1);
   255 qed "real_add_minus_eq_minus";
   256 
   257 Goal "? (y::real). x = -y";
   258 by (cut_inst_tac [("x","x")] real_minus_ex 1);
   259 by (etac exE 1 THEN dtac real_add_minus_eq_minus 1);
   260 by (Fast_tac 1);
   261 qed "real_as_add_inverse_ex";
   262 
   263 Goal "-(x + y) = (-x) + (- y :: real)";
   264 by (res_inst_tac [("z","x")] eq_Abs_real 1);
   265 by (res_inst_tac [("z","y")] eq_Abs_real 1);
   266 by (auto_tac (claset(),simpset() addsimps [real_minus,real_add]));
   267 qed "real_minus_add_distrib";
   268 
   269 Addsimps [real_minus_add_distrib];
   270 
   271 Goal "((x::real) + y = x + z) = (y = z)";
   272 by (Step_tac 1);
   273 by (dres_inst_tac [("f","%t. (-x) + t")] arg_cong 1);
   274 by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
   275 qed "real_add_left_cancel";
   276 
   277 Goal "(y + (x::real)= z + x) = (y = z)";
   278 by (simp_tac (simpset() addsimps [real_add_commute,real_add_left_cancel]) 1);
   279 qed "real_add_right_cancel";
   280 
   281 Goal "((x::real) = y) = (0r = x + (- y))";
   282 by (Step_tac 1);
   283 by (res_inst_tac [("x1","-y")] 
   284       (real_add_right_cancel RS iffD1) 2);
   285 by Auto_tac;
   286 qed "real_eq_minus_iff"; 
   287 
   288 Goal "((x::real) = y) = (x + (- y) = 0r)";
   289 by (Step_tac 1);
   290 by (res_inst_tac [("x1","-y")] 
   291       (real_add_right_cancel RS iffD1) 2);
   292 by Auto_tac;
   293 qed "real_eq_minus_iff2"; 
   294 
   295 Goal "0r - x = -x";
   296 by (simp_tac (simpset() addsimps [real_diff_def]) 1);
   297 qed "real_diff_0";
   298 
   299 Goal "x - 0r = x";
   300 by (simp_tac (simpset() addsimps [real_diff_def]) 1);
   301 qed "real_diff_0_right";
   302 
   303 Goal "x - x = 0r";
   304 by (simp_tac (simpset() addsimps [real_diff_def]) 1);
   305 qed "real_diff_self";
   306 
   307 Addsimps [real_diff_0, real_diff_0_right, real_diff_self];
   308 
   309 
   310 (*** Congruence property for multiplication ***)
   311 
   312 Goal "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> \
   313 \         x * x1 + y * y1 + (x * y2 + x2 * y) = \
   314 \         x * x2 + y * y2 + (x * y1 + x1 * y)";
   315 by (asm_full_simp_tac (simpset() addsimps [preal_add_left_commute,
   316     preal_add_assoc RS sym,preal_add_mult_distrib2 RS sym]) 1);
   317 by (rtac (preal_mult_commute RS subst) 1);
   318 by (res_inst_tac [("y1","x2")] (preal_mult_commute RS subst) 1);
   319 by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc,
   320     preal_add_mult_distrib2 RS sym]) 1);
   321 by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1);
   322 qed "real_mult_congruent2_lemma";
   323 
   324 Goal 
   325     "congruent2 realrel (%p1 p2.                  \
   326 \         split (%x1 y1. split (%x2 y2. realrel^^{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)";
   327 by (rtac (equiv_realrel RS congruent2_commuteI) 1);
   328 by Safe_tac;
   329 by (rewtac split_def);
   330 by (asm_simp_tac (simpset() addsimps [preal_mult_commute,preal_add_commute]) 1);
   331 by (auto_tac (claset(),simpset() addsimps [real_mult_congruent2_lemma]));
   332 qed "real_mult_congruent2";
   333 
   334 (*Resolve th against the corresponding facts for real_mult*)
   335 val real_mult_ize = RSLIST [equiv_realrel, real_mult_congruent2];
   336 
   337 Goalw [real_mult_def]
   338    "Abs_real((realrel^^{(x1,y1)})) * Abs_real((realrel^^{(x2,y2)})) =   \
   339 \   Abs_real(realrel ^^ {(x1*x2+y1*y2,x1*y2+x2*y1)})";
   340 by (simp_tac (simpset() addsimps [real_mult_ize UN_equiv_class2]) 1);
   341 qed "real_mult";
   342 
   343 Goal "(z::real) * w = w * z";
   344 by (res_inst_tac [("z","z")] eq_Abs_real 1);
   345 by (res_inst_tac [("z","w")] eq_Abs_real 1);
   346 by (asm_simp_tac
   347     (simpset() addsimps [real_mult] @ preal_add_ac @ preal_mult_ac) 1);
   348 qed "real_mult_commute";
   349 
   350 Goal "((z1::real) * z2) * z3 = z1 * (z2 * z3)";
   351 by (res_inst_tac [("z","z1")] eq_Abs_real 1);
   352 by (res_inst_tac [("z","z2")] eq_Abs_real 1);
   353 by (res_inst_tac [("z","z3")] eq_Abs_real 1);
   354 by (asm_simp_tac (simpset() addsimps [preal_add_mult_distrib2,real_mult] @ 
   355                                      preal_add_ac @ preal_mult_ac) 1);
   356 qed "real_mult_assoc";
   357 
   358 qed_goal "real_mult_left_commute" thy
   359     "(z1::real) * (z2 * z3) = z2 * (z1 * z3)"
   360  (fn _ => [rtac (real_mult_commute RS trans) 1, rtac (real_mult_assoc RS trans) 1,
   361            rtac (real_mult_commute RS arg_cong) 1]);
   362 
   363 (* real multiplication is an AC operator *)
   364 bind_thms ("real_mult_ac", [real_mult_assoc, real_mult_commute, real_mult_left_commute]);
   365 
   366 Goalw [real_one_def,pnat_one_def] "1r * z = z";
   367 by (res_inst_tac [("z","z")] eq_Abs_real 1);
   368 by (asm_full_simp_tac
   369     (simpset() addsimps [real_mult,
   370 			 preal_add_mult_distrib2,preal_mult_1_right] 
   371                         @ preal_mult_ac @ preal_add_ac) 1);
   372 qed "real_mult_1";
   373 
   374 Addsimps [real_mult_1];
   375 
   376 Goal "z * 1r = z";
   377 by (simp_tac (simpset() addsimps [real_mult_commute]) 1);
   378 qed "real_mult_1_right";
   379 
   380 Addsimps [real_mult_1_right];
   381 
   382 Goalw [real_zero_def,pnat_one_def] "0r * z = 0r";
   383 by (res_inst_tac [("z","z")] eq_Abs_real 1);
   384 by (asm_full_simp_tac (simpset() addsimps [real_mult,
   385     preal_add_mult_distrib2,preal_mult_1_right] 
   386     @ preal_mult_ac @ preal_add_ac) 1);
   387 qed "real_mult_0";
   388 
   389 Goal "z * 0r = 0r";
   390 by (simp_tac (simpset() addsimps [real_mult_commute, real_mult_0]) 1);
   391 qed "real_mult_0_right";
   392 
   393 Addsimps [real_mult_0_right, real_mult_0];
   394 
   395 Goal "-(x * y) = (-x) * (y::real)";
   396 by (res_inst_tac [("z","x")] eq_Abs_real 1);
   397 by (res_inst_tac [("z","y")] eq_Abs_real 1);
   398 by (auto_tac (claset(),
   399 	      simpset() addsimps [real_minus,real_mult] 
   400     @ preal_mult_ac @ preal_add_ac));
   401 qed "real_minus_mult_eq1";
   402 
   403 Goal "-(x * y) = x * (- y :: real)";
   404 by (res_inst_tac [("z","x")] eq_Abs_real 1);
   405 by (res_inst_tac [("z","y")] eq_Abs_real 1);
   406 by (auto_tac (claset(),
   407 	      simpset() addsimps [real_minus,real_mult] 
   408     @ preal_mult_ac @ preal_add_ac));
   409 qed "real_minus_mult_eq2";
   410 
   411 Goal "(- 1r) * z = -z";
   412 by (simp_tac (simpset() addsimps [real_minus_mult_eq1 RS sym]) 1);
   413 qed "real_mult_minus_1";
   414 
   415 Addsimps [real_mult_minus_1];
   416 
   417 Goal "z * (- 1r) = -z";
   418 by (stac real_mult_commute 1);
   419 by (Simp_tac 1);
   420 qed "real_mult_minus_1_right";
   421 
   422 Addsimps [real_mult_minus_1_right];
   423 
   424 Goal "(-x) * (-y) = x * (y::real)";
   425 by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym,
   426 				       real_minus_mult_eq1 RS sym]) 1);
   427 qed "real_minus_mult_cancel";
   428 
   429 Addsimps [real_minus_mult_cancel];
   430 
   431 Goal "(-x) * y = x * (- y :: real)";
   432 by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym,
   433 				       real_minus_mult_eq1 RS sym]) 1);
   434 qed "real_minus_mult_commute";
   435 
   436 (*-----------------------------------------------------------------------------
   437 
   438  ----------------------------------------------------------------------------*)
   439 
   440 (** Lemmas **)
   441 
   442 qed_goal "real_add_assoc_cong" thy
   443     "!!z. (z::real) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
   444  (fn _ => [(asm_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1)]);
   445 
   446 qed_goal "real_add_assoc_swap" thy "(z::real) + (v + w) = v + (z + w)"
   447  (fn _ => [(REPEAT (ares_tac [real_add_commute RS real_add_assoc_cong] 1))]);
   448 
   449 Goal "((z1::real) + z2) * w = (z1 * w) + (z2 * w)";
   450 by (res_inst_tac [("z","z1")] eq_Abs_real 1);
   451 by (res_inst_tac [("z","z2")] eq_Abs_real 1);
   452 by (res_inst_tac [("z","w")] eq_Abs_real 1);
   453 by (asm_simp_tac 
   454     (simpset() addsimps [preal_add_mult_distrib2, real_add, real_mult] @ 
   455                         preal_add_ac @ preal_mult_ac) 1);
   456 qed "real_add_mult_distrib";
   457 
   458 val real_mult_commute'= read_instantiate [("z","w")] real_mult_commute;
   459 
   460 Goal "(w::real) * (z1 + z2) = (w * z1) + (w * z2)";
   461 by (simp_tac (simpset() addsimps [real_mult_commute',real_add_mult_distrib]) 1);
   462 qed "real_add_mult_distrib2";
   463 
   464 Goalw [real_diff_def] "((z1::real) - z2) * w = (z1 * w) - (z2 * w)";
   465 by (simp_tac (simpset() addsimps [real_add_mult_distrib, 
   466 				  real_minus_mult_eq1]) 1);
   467 qed "real_diff_mult_distrib";
   468 
   469 Goal "(w::real) * (z1 - z2) = (w * z1) - (w * z2)";
   470 by (simp_tac (simpset() addsimps [real_mult_commute', 
   471 				  real_diff_mult_distrib]) 1);
   472 qed "real_diff_mult_distrib2";
   473 
   474 (*** one and zero are distinct ***)
   475 Goalw [real_zero_def,real_one_def] "0r ~= 1r";
   476 by (auto_tac (claset(),
   477          simpset() addsimps [preal_self_less_add_left RS preal_not_refl2]));
   478 qed "real_zero_not_eq_one";
   479 
   480 (*** existence of inverse ***)
   481 (** lemma -- alternative definition for 0r **)
   482 Goalw [real_zero_def] "0r = Abs_real (realrel ^^ {(x, x)})";
   483 by (auto_tac (claset(),simpset() addsimps [preal_add_commute]));
   484 qed "real_zero_iff";
   485 
   486 Goalw [real_zero_def,real_one_def] 
   487           "!!(x::real). x ~= 0r ==> ? y. x*y = 1r";
   488 by (res_inst_tac [("z","x")] eq_Abs_real 1);
   489 by (cut_inst_tac [("r1.0","xa"),("r2.0","y")] preal_linear 1);
   490 by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
   491            simpset() addsimps [real_zero_iff RS sym]));
   492 by (res_inst_tac [("x","Abs_real (realrel ^^ \
   493 \   {(preal_of_prat(prat_of_pnat 1p),pinv(D)+\
   494 \    preal_of_prat(prat_of_pnat 1p))})")] exI 1);
   495 by (res_inst_tac [("x","Abs_real (realrel ^^ \
   496 \   {(pinv(D)+preal_of_prat(prat_of_pnat 1p),\
   497 \    preal_of_prat(prat_of_pnat 1p))})")] exI 2);
   498 by (auto_tac (claset(),
   499 	      simpset() addsimps [real_mult,
   500     pnat_one_def,preal_mult_1_right,preal_add_mult_distrib2,
   501     preal_add_mult_distrib,preal_mult_1,preal_mult_inv_right] 
   502     @ preal_add_ac @ preal_mult_ac));
   503 qed "real_mult_inv_right_ex";
   504 
   505 Goal "!!(x::real). x ~= 0r ==> ? y. y*x = 1r";
   506 by (asm_simp_tac (simpset() addsimps [real_mult_commute,
   507 				      real_mult_inv_right_ex]) 1);
   508 qed "real_mult_inv_left_ex";
   509 
   510 Goalw [rinv_def] "x ~= 0r ==> rinv(x)*x = 1r";
   511 by (ftac real_mult_inv_left_ex 1);
   512 by (Step_tac 1);
   513 by (rtac selectI2 1);
   514 by Auto_tac;
   515 qed "real_mult_inv_left";
   516 
   517 Goal "x ~= 0r ==> x*rinv(x) = 1r";
   518 by (auto_tac (claset() addIs [real_mult_commute RS subst],
   519               simpset() addsimps [real_mult_inv_left]));
   520 qed "real_mult_inv_right";
   521 
   522 Goal "(c::real) ~= 0r ==> (c*a=c*b) = (a=b)";
   523 by Auto_tac;
   524 by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
   525 by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
   526 qed "real_mult_left_cancel";
   527     
   528 Goal "(c::real) ~= 0r ==> (a*c=b*c) = (a=b)";
   529 by (Step_tac 1);
   530 by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
   531 by (asm_full_simp_tac
   532     (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
   533 qed "real_mult_right_cancel";
   534 
   535 Goal "c*a ~= c*b ==> a ~= b";
   536 by Auto_tac;
   537 qed "real_mult_left_cancel_ccontr";
   538 
   539 Goal "a*c ~= b*c ==> a ~= b";
   540 by Auto_tac;
   541 qed "real_mult_right_cancel_ccontr";
   542 
   543 Goalw [rinv_def] "x ~= 0r ==> rinv(x) ~= 0r";
   544 by (ftac real_mult_inv_left_ex 1);
   545 by (etac exE 1);
   546 by (rtac selectI2 1);
   547 by (auto_tac (claset(),
   548 	      simpset() addsimps [real_mult_0,
   549     real_zero_not_eq_one]));
   550 qed "rinv_not_zero";
   551 
   552 Addsimps [real_mult_inv_left,real_mult_inv_right];
   553 
   554 Goal "[| x ~= 0r; y ~= 0r |] ==> x * y ~= 0r";
   555 by (Step_tac 1);
   556 by (dres_inst_tac [("f","%z. rinv x*z")] arg_cong 1);
   557 by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
   558 qed "real_mult_not_zero";
   559 
   560 bind_thm ("real_mult_not_zeroE",real_mult_not_zero RS notE);
   561 
   562 Goal "x ~= 0r ==> rinv(rinv x) = x";
   563 by (res_inst_tac [("c1","rinv x")] (real_mult_right_cancel RS iffD1) 1);
   564 by (etac rinv_not_zero 1);
   565 by (auto_tac (claset() addDs [rinv_not_zero],simpset()));
   566 qed "real_rinv_rinv";
   567 
   568 Goalw [rinv_def] "rinv(1r) = 1r";
   569 by (cut_facts_tac [real_zero_not_eq_one RS 
   570        not_sym RS real_mult_inv_left_ex] 1);
   571 by (etac exE 1);
   572 by (rtac selectI2 1);
   573 by (auto_tac (claset(),
   574 	      simpset() addsimps 
   575     [real_zero_not_eq_one RS not_sym]));
   576 qed "real_rinv_1";
   577 Addsimps [real_rinv_1];
   578 
   579 Goal "x ~= 0r ==> rinv(-x) = -rinv(x)";
   580 by (res_inst_tac [("c1","-x")] (real_mult_right_cancel RS iffD1) 1);
   581 by Auto_tac;
   582 qed "real_minus_rinv";
   583 
   584 Goal "[| x ~= 0r; y ~= 0r |] ==> rinv(x*y) = rinv(x)*rinv(y)";
   585 by (forw_inst_tac [("y","y")] real_mult_not_zero 1 THEN assume_tac 1);
   586 by (res_inst_tac [("c1","x")] (real_mult_left_cancel RS iffD1) 1);
   587 by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]));
   588 by (res_inst_tac [("c1","y")] (real_mult_left_cancel RS iffD1) 1);
   589 by (auto_tac (claset(),simpset() addsimps [real_mult_left_commute]));
   590 by (asm_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
   591 qed "real_rinv_distrib";
   592 
   593 (*---------------------------------------------------------
   594      Theorems for ordering
   595  --------------------------------------------------------*)
   596 (* prove introduction and elimination rules for real_less *)
   597 
   598 (* real_less is a strong order i.e. nonreflexive and transitive *)
   599 
   600 (*** lemmas ***)
   601 Goal "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y";
   602 by (asm_simp_tac (simpset() addsimps [preal_add_commute]) 1);
   603 qed "preal_lemma_eq_rev_sum";
   604 
   605 Goal "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1";
   606 by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
   607 qed "preal_add_left_commute_cancel";
   608 
   609 Goal "!!(x::preal). [| x + y2a = x2a + y; \
   610 \                      x + y2b = x2b + y |] \
   611 \                   ==> x2a + y2b = x2b + y2a";
   612 by (dtac preal_lemma_eq_rev_sum 1);
   613 by (assume_tac 1);
   614 by (thin_tac "x + y2b = x2b + y" 1);
   615 by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
   616 by (dtac preal_add_left_commute_cancel 1);
   617 by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
   618 qed "preal_lemma_for_not_refl";
   619 
   620 Goal "~ (R::real) < R";
   621 by (res_inst_tac [("z","R")] eq_Abs_real 1);
   622 by (auto_tac (claset(),simpset() addsimps [real_less_def]));
   623 by (dtac preal_lemma_for_not_refl 1);
   624 by (assume_tac 1 THEN rotate_tac 2 1);
   625 by (auto_tac (claset(),simpset() addsimps [preal_less_not_refl]));
   626 qed "real_less_not_refl";
   627 
   628 (*** y < y ==> P ***)
   629 bind_thm("real_less_irrefl", real_less_not_refl RS notE);
   630 AddSEs [real_less_irrefl];
   631 
   632 Goal "!!(x::real). x < y ==> x ~= y";
   633 by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
   634 qed "real_not_refl2";
   635 
   636 (* lemma re-arranging and eliminating terms *)
   637 Goal "!! (a::preal). [| a + b = c + d; \
   638 \            x2b + d + (c + y2e) < a + y2b + (x2e + b) |] \
   639 \         ==> x2b + y2e < x2e + y2b";
   640 by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
   641 by (res_inst_tac [("C","c+d")] preal_add_left_less_cancel 1);
   642 by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
   643 qed "preal_lemma_trans";
   644 
   645 (** heavy re-writing involved*)
   646 Goal "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
   647 by (res_inst_tac [("z","R1")] eq_Abs_real 1);
   648 by (res_inst_tac [("z","R2")] eq_Abs_real 1);
   649 by (res_inst_tac [("z","R3")] eq_Abs_real 1);
   650 by (auto_tac (claset(),simpset() addsimps [real_less_def]));
   651 by (REPEAT(rtac exI 1));
   652 by (EVERY[rtac conjI 1, rtac conjI 2]);
   653 by (REPEAT(Blast_tac 2));
   654 by (dtac preal_lemma_for_not_refl 1 THEN assume_tac 1);
   655 by (blast_tac (claset() addDs [preal_add_less_mono] 
   656     addIs [preal_lemma_trans]) 1);
   657 qed "real_less_trans";
   658 
   659 Goal "!! (R1::real). [| R1 < R2; R2 < R1 |] ==> P";
   660 by (dtac real_less_trans 1 THEN assume_tac 1);
   661 by (asm_full_simp_tac (simpset() addsimps [real_less_not_refl]) 1);
   662 qed "real_less_asym";
   663 
   664 (****)(****)(****)(****)(****)(****)(****)(****)(****)(****)
   665     (****** Map and more real_less ******)
   666 (*** mapping from preal into real ***)
   667 Goalw [real_of_preal_def] 
   668      "real_of_preal ((z1::preal) + z2) = \
   669 \     real_of_preal z1 + real_of_preal z2";
   670 by (asm_simp_tac (simpset() addsimps [real_add,
   671        preal_add_mult_distrib,preal_mult_1] addsimps preal_add_ac) 1);
   672 qed "real_of_preal_add";
   673 
   674 Goalw [real_of_preal_def] 
   675      "real_of_preal ((z1::preal) * z2) = \
   676 \     real_of_preal z1* real_of_preal z2";
   677 by (full_simp_tac (simpset() addsimps [real_mult,
   678         preal_add_mult_distrib2,preal_mult_1,
   679         preal_mult_1_right,pnat_one_def] 
   680         @ preal_add_ac @ preal_mult_ac) 1);
   681 qed "real_of_preal_mult";
   682 
   683 Goalw [real_of_preal_def]
   684       "!!(x::preal). y < x ==> \
   685 \      ? m. Abs_real (realrel ^^ {(x,y)}) = real_of_preal m";
   686 by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
   687     simpset() addsimps preal_add_ac));
   688 qed "real_of_preal_ExI";
   689 
   690 Goalw [real_of_preal_def]
   691       "!!(x::preal). ? m. Abs_real (realrel ^^ {(x,y)}) = \
   692 \                    real_of_preal m ==> y < x";
   693 by (auto_tac (claset(),
   694 	      simpset() addsimps 
   695     [preal_add_commute,preal_add_assoc]));
   696 by (asm_full_simp_tac (simpset() addsimps 
   697     [preal_add_assoc RS sym,preal_self_less_add_left]) 1);
   698 qed "real_of_preal_ExD";
   699 
   700 Goal "(? m. Abs_real (realrel ^^ {(x,y)}) = real_of_preal m) = (y < x)";
   701 by (blast_tac (claset() addSIs [real_of_preal_ExI,real_of_preal_ExD]) 1);
   702 qed "real_of_preal_iff";
   703 
   704 (*** Gleason prop 9-4.4 p 127 ***)
   705 Goalw [real_of_preal_def,real_zero_def] 
   706       "? m. (x::real) = real_of_preal m | x = 0r | x = -(real_of_preal m)";
   707 by (res_inst_tac [("z","x")] eq_Abs_real 1);
   708 by (auto_tac (claset(),simpset() addsimps [real_minus] @ preal_add_ac));
   709 by (cut_inst_tac [("r1.0","x"),("r2.0","y")] preal_linear 1);
   710 by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
   711     simpset() addsimps [preal_add_assoc RS sym]));
   712 by (auto_tac (claset(),simpset() addsimps [preal_add_commute]));
   713 qed "real_of_preal_trichotomy";
   714 
   715 Goal "!!P. [| !!m. x = real_of_preal m ==> P; \
   716 \             x = 0r ==> P; \
   717 \             !!m. x = -(real_of_preal m) ==> P |] ==> P";
   718 by (cut_inst_tac [("x","x")] real_of_preal_trichotomy 1);
   719 by Auto_tac;
   720 qed "real_of_preal_trichotomyE";
   721 
   722 Goalw [real_of_preal_def] 
   723       "real_of_preal m1 < real_of_preal m2 ==> m1 < m2";
   724 by (auto_tac (claset(),simpset() addsimps [real_less_def] @ preal_add_ac));
   725 by (auto_tac (claset(),simpset() addsimps [preal_add_assoc RS sym]));
   726 by (auto_tac (claset(),simpset() addsimps preal_add_ac));
   727 qed "real_of_preal_lessD";
   728 
   729 Goal "m1 < m2 ==> real_of_preal m1 < real_of_preal m2";
   730 by (dtac preal_less_add_left_Ex 1);
   731 by (auto_tac (claset(),
   732 	      simpset() addsimps [real_of_preal_add,
   733     real_of_preal_def,real_less_def]));
   734 by (REPEAT(rtac exI 1));
   735 by (EVERY[rtac conjI 1, rtac conjI 2]);
   736 by (REPEAT(Blast_tac 2));
   737 by (simp_tac (simpset() addsimps [preal_self_less_add_left] 
   738     delsimps [preal_add_less_iff2]) 1);
   739 qed "real_of_preal_lessI";
   740 
   741 Goal "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)";
   742 by (blast_tac (claset() addIs [real_of_preal_lessI,real_of_preal_lessD]) 1);
   743 qed "real_of_preal_less_iff1";
   744 
   745 Addsimps [real_of_preal_less_iff1];
   746 
   747 Goal "- real_of_preal m < real_of_preal m";
   748 by (auto_tac (claset(),
   749 	      simpset() addsimps 
   750     [real_of_preal_def,real_less_def,real_minus]));
   751 by (REPEAT(rtac exI 1));
   752 by (EVERY[rtac conjI 1, rtac conjI 2]);
   753 by (REPEAT(Blast_tac 2));
   754 by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
   755 by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
   756     preal_add_assoc RS sym]) 1);
   757 qed "real_of_preal_minus_less_self";
   758 
   759 Goalw [real_zero_def] "- real_of_preal m < 0r";
   760 by (auto_tac (claset(),
   761 	      simpset() addsimps [real_of_preal_def,
   762 				  real_less_def,real_minus]));
   763 by (REPEAT(rtac exI 1));
   764 by (EVERY[rtac conjI 1, rtac conjI 2]);
   765 by (REPEAT(Blast_tac 2));
   766 by (full_simp_tac (simpset() addsimps 
   767   [preal_self_less_add_right] @ preal_add_ac) 1);
   768 qed "real_of_preal_minus_less_zero";
   769 
   770 Goal "~ 0r < - real_of_preal m";
   771 by (cut_facts_tac [real_of_preal_minus_less_zero] 1);
   772 by (fast_tac (claset() addDs [real_less_trans] 
   773                         addEs [real_less_irrefl]) 1);
   774 qed "real_of_preal_not_minus_gt_zero";
   775 
   776 Goalw [real_zero_def] "0r < real_of_preal m";
   777 by (auto_tac (claset(),simpset() addsimps 
   778    [real_of_preal_def,real_less_def,real_minus]));
   779 by (REPEAT(rtac exI 1));
   780 by (EVERY[rtac conjI 1, rtac conjI 2]);
   781 by (REPEAT(Blast_tac 2));
   782 by (full_simp_tac (simpset() addsimps 
   783 		   [preal_self_less_add_right] @ preal_add_ac) 1);
   784 qed "real_of_preal_zero_less";
   785 
   786 Goal "~ real_of_preal m < 0r";
   787 by (cut_facts_tac [real_of_preal_zero_less] 1);
   788 by (blast_tac (claset() addDs [real_less_trans] 
   789                         addEs [real_less_irrefl]) 1);
   790 qed "real_of_preal_not_less_zero";
   791 
   792 Goal "0r < - (- real_of_preal m)";
   793 by (simp_tac (simpset() addsimps 
   794     [real_of_preal_zero_less]) 1);
   795 qed "real_minus_minus_zero_less";
   796 
   797 (* another lemma *)
   798 Goalw [real_zero_def]
   799       "0r < real_of_preal m + real_of_preal m1";
   800 by (auto_tac (claset(),
   801 	      simpset() addsimps [real_of_preal_def,
   802 				  real_less_def,real_add]));
   803 by (REPEAT(rtac exI 1));
   804 by (EVERY[rtac conjI 1, rtac conjI 2]);
   805 by (REPEAT(Blast_tac 2));
   806 by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
   807 by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
   808     preal_add_assoc RS sym]) 1);
   809 qed "real_of_preal_sum_zero_less";
   810 
   811 Goal "- real_of_preal m < real_of_preal m1";
   812 by (auto_tac (claset(),
   813 	      simpset() addsimps [real_of_preal_def,
   814               real_less_def,real_minus]));
   815 by (REPEAT(rtac exI 1));
   816 by (EVERY[rtac conjI 1, rtac conjI 2]);
   817 by (REPEAT(Blast_tac 2));
   818 by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
   819 by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
   820     preal_add_assoc RS sym]) 1);
   821 qed "real_of_preal_minus_less_all";
   822 
   823 Goal "~ real_of_preal m < - real_of_preal m1";
   824 by (cut_facts_tac [real_of_preal_minus_less_all] 1);
   825 by (blast_tac (claset() addDs [real_less_trans] 
   826                addEs [real_less_irrefl]) 1);
   827 qed "real_of_preal_not_minus_gt_all";
   828 
   829 Goal "- real_of_preal m1 < - real_of_preal m2 \
   830 \     ==> real_of_preal m2 < real_of_preal m1";
   831 by (auto_tac (claset(),
   832 	      simpset() addsimps [real_of_preal_def,
   833               real_less_def,real_minus]));
   834 by (REPEAT(rtac exI 1));
   835 by (EVERY[rtac conjI 1, rtac conjI 2]);
   836 by (REPEAT(Blast_tac 2));
   837 by (auto_tac (claset(),simpset() addsimps preal_add_ac));
   838 by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
   839 by (auto_tac (claset(),simpset() addsimps preal_add_ac));
   840 qed "real_of_preal_minus_less_rev1";
   841 
   842 Goal "real_of_preal m1 < real_of_preal m2 \
   843 \     ==> - real_of_preal m2 < - real_of_preal m1";
   844 by (auto_tac (claset(),
   845 	      simpset() addsimps [real_of_preal_def,
   846               real_less_def,real_minus]));
   847 by (REPEAT(rtac exI 1));
   848 by (EVERY[rtac conjI 1, rtac conjI 2]);
   849 by (REPEAT(Blast_tac 2));
   850 by (auto_tac (claset(),simpset() addsimps preal_add_ac));
   851 by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
   852 by (auto_tac (claset(),simpset() addsimps preal_add_ac));
   853 qed "real_of_preal_minus_less_rev2";
   854 
   855 Goal "(- real_of_preal m1 < - real_of_preal m2) = \
   856 \     (real_of_preal m2 < real_of_preal m1)";
   857 by (blast_tac (claset() addSIs [real_of_preal_minus_less_rev1,
   858                real_of_preal_minus_less_rev2]) 1);
   859 qed "real_of_preal_minus_less_rev_iff";
   860 
   861 Addsimps [real_of_preal_minus_less_rev_iff];
   862 
   863 (*** linearity ***)
   864 Goal "(R1::real) < R2 | R1 = R2 | R2 < R1";
   865 by (res_inst_tac [("x","R1")]  real_of_preal_trichotomyE 1);
   866 by (ALLGOALS(res_inst_tac [("x","R2")]  real_of_preal_trichotomyE));
   867 by (auto_tac (claset() addSDs [preal_le_anti_sym],
   868               simpset() addsimps [preal_less_le_iff,real_of_preal_minus_less_zero,
   869                real_of_preal_zero_less,real_of_preal_minus_less_all]));
   870 qed "real_linear";
   871 
   872 Goal "!!w::real. (w ~= z) = (w<z | z<w)";
   873 by (cut_facts_tac [real_linear] 1);
   874 by (Blast_tac 1);
   875 qed "real_neq_iff";
   876 
   877 Goal "!!(R1::real). [| R1 < R2 ==> P;  R1 = R2 ==> P; \
   878 \                      R2 < R1 ==> P |] ==> P";
   879 by (cut_inst_tac [("R1.0","R1"),("R2.0","R2")] real_linear 1);
   880 by Auto_tac;
   881 qed "real_linear_less2";
   882 
   883 (*** Properties of <= ***)
   884 
   885 Goalw [real_le_def] "~(w < z) ==> z <= (w::real)";
   886 by (assume_tac 1);
   887 qed "real_leI";
   888 
   889 Goalw [real_le_def] "z<=w ==> ~(w<(z::real))";
   890 by (assume_tac 1);
   891 qed "real_leD";
   892 
   893 bind_thm ("real_leE", make_elim real_leD);
   894 
   895 Goal "(~(w < z)) = (z <= (w::real))";
   896 by (blast_tac (claset() addSIs [real_leI,real_leD]) 1);
   897 qed "real_less_le_iff";
   898 
   899 Goalw [real_le_def] "~ z <= w ==> w<(z::real)";
   900 by (Blast_tac 1);
   901 qed "not_real_leE";
   902 
   903 Goalw [real_le_def] "z < w ==> z <= (w::real)";
   904 by (blast_tac (claset() addEs [real_less_asym]) 1);
   905 qed "real_less_imp_le";
   906 
   907 Goalw [real_le_def] "!!(x::real). x <= y ==> x < y | x = y";
   908 by (cut_facts_tac [real_linear] 1);
   909 by (blast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1);
   910 qed "real_le_imp_less_or_eq";
   911 
   912 Goalw [real_le_def] "z<w | z=w ==> z <=(w::real)";
   913 by (cut_facts_tac [real_linear] 1);
   914 by (fast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1);
   915 qed "real_less_or_eq_imp_le";
   916 
   917 Goal "(x <= (y::real)) = (x < y | x=y)";
   918 by (REPEAT(ares_tac [iffI, real_less_or_eq_imp_le, real_le_imp_less_or_eq] 1));
   919 qed "real_le_less";
   920 
   921 Goal "w <= (w::real)";
   922 by (simp_tac (simpset() addsimps [real_le_less]) 1);
   923 qed "real_le_refl";
   924 
   925 AddIffs [real_le_refl];
   926 
   927 (* Axiom 'linorder_linear' of class 'linorder': *)
   928 Goal "(z::real) <= w | w <= z";
   929 by (simp_tac (simpset() addsimps [real_le_less]) 1);
   930 by (cut_facts_tac [real_linear] 1);
   931 by (Blast_tac 1);
   932 qed "real_le_linear";
   933 
   934 Goal "[| i <= j; j < k |] ==> i < (k::real)";
   935 by (dtac real_le_imp_less_or_eq 1);
   936 by (blast_tac (claset() addIs [real_less_trans]) 1);
   937 qed "real_le_less_trans";
   938 
   939 Goal "!! (i::real). [| i < j; j <= k |] ==> i < k";
   940 by (dtac real_le_imp_less_or_eq 1);
   941 by (blast_tac (claset() addIs [real_less_trans]) 1);
   942 qed "real_less_le_trans";
   943 
   944 Goal "[| i <= j; j <= k |] ==> i <= (k::real)";
   945 by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq,
   946             rtac real_less_or_eq_imp_le, blast_tac (claset() addIs [real_less_trans])]);
   947 qed "real_le_trans";
   948 
   949 Goal "[| z <= w; w <= z |] ==> z = (w::real)";
   950 by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq,
   951             fast_tac (claset() addEs [real_less_irrefl,real_less_asym])]);
   952 qed "real_le_anti_sym";
   953 
   954 Goal "[| ~ y < x; y ~= x |] ==> x < (y::real)";
   955 by (rtac not_real_leE 1);
   956 by (blast_tac (claset() addDs [real_le_imp_less_or_eq]) 1);
   957 qed "not_less_not_eq_real_less";
   958 
   959 (* Axiom 'order_less_le' of class 'order': *)
   960 Goal "(w::real) < z = (w <= z & w ~= z)";
   961 by (simp_tac (simpset() addsimps [real_le_def, real_neq_iff]) 1);
   962 by (blast_tac (claset() addSEs [real_less_asym]) 1);
   963 qed "real_less_le";
   964 
   965 Goal "(0r < -R) = (R < 0r)";
   966 by (res_inst_tac [("x","R")]  real_of_preal_trichotomyE 1);
   967 by (auto_tac (claset(),
   968 	      simpset() addsimps [real_of_preal_not_minus_gt_zero,
   969                         real_of_preal_not_less_zero,real_of_preal_zero_less,
   970                         real_of_preal_minus_less_zero]));
   971 qed "real_minus_zero_less_iff";
   972 
   973 Addsimps [real_minus_zero_less_iff];
   974 
   975 Goal "(-R < 0r) = (0r < R)";
   976 by (res_inst_tac [("x","R")]  real_of_preal_trichotomyE 1);
   977 by (auto_tac (claset(),
   978 	      simpset() addsimps [real_of_preal_not_minus_gt_zero,
   979                         real_of_preal_not_less_zero,real_of_preal_zero_less,
   980                         real_of_preal_minus_less_zero]));
   981 qed "real_minus_zero_less_iff2";
   982 
   983 (*Alternative definition for real_less*)
   984 Goal "R < S ==> ? T. 0r < T & R + T = S";
   985 by (res_inst_tac [("x","R")]  real_of_preal_trichotomyE 1);
   986 by (ALLGOALS(res_inst_tac [("x","S")]  real_of_preal_trichotomyE));
   987 by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
   988 	      simpset() addsimps [real_of_preal_not_minus_gt_all,
   989 				  real_of_preal_add, real_of_preal_not_less_zero,
   990 				  real_less_not_refl,
   991 				  real_of_preal_not_minus_gt_zero]));
   992 by (res_inst_tac [("x","real_of_preal D")] exI 1);
   993 by (res_inst_tac [("x","real_of_preal m+real_of_preal ma")] exI 2);
   994 by (res_inst_tac [("x","real_of_preal m")] exI 3);
   995 by (res_inst_tac [("x","real_of_preal D")] exI 4);
   996 by (auto_tac (claset(),
   997 	      simpset() addsimps [real_of_preal_zero_less,
   998 				  real_of_preal_sum_zero_less,real_add_assoc]));
   999 qed "real_less_add_positive_left_Ex";
  1000 
  1001 (** change naff name(s)! **)
  1002 Goal "(W < S) ==> (0r < S + (-W))";
  1003 by (dtac real_less_add_positive_left_Ex 1);
  1004 by (auto_tac (claset(),
  1005 	      simpset() addsimps [real_add_minus,
  1006     real_add_zero_right] @ real_add_ac));
  1007 qed "real_less_sum_gt_zero";
  1008 
  1009 Goal "!!S::real. T = S + W ==> S = T + (-W)";
  1010 by (asm_simp_tac (simpset() addsimps real_add_ac) 1);
  1011 qed "real_lemma_change_eq_subj";
  1012 
  1013 (* FIXME: long! *)
  1014 Goal "(0r < S + (-W)) ==> (W < S)";
  1015 by (rtac ccontr 1);
  1016 by (dtac (real_leI RS real_le_imp_less_or_eq) 1);
  1017 by (auto_tac (claset(),
  1018 	      simpset() addsimps [real_less_not_refl]));
  1019 by (EVERY1[dtac real_less_add_positive_left_Ex, etac exE, etac conjE]);
  1020 by (Asm_full_simp_tac 1);
  1021 by (dtac real_lemma_change_eq_subj 1);
  1022 by Auto_tac;
  1023 by (dtac real_less_sum_gt_zero 1);
  1024 by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1);
  1025 by (EVERY1[rotate_tac 1, dtac (real_add_left_commute RS ssubst)]);
  1026 by (auto_tac (claset() addEs [real_less_asym], simpset()));
  1027 qed "real_sum_gt_zero_less";
  1028 
  1029 Goal "(0r < S + (-W)) = (W < S)";
  1030 by (blast_tac (claset() addIs [real_less_sum_gt_zero,
  1031 			       real_sum_gt_zero_less]) 1);
  1032 qed "real_less_sum_gt_0_iff";
  1033 
  1034 
  1035 Goalw [real_diff_def] "(x<y) = (x-y < 0r)";
  1036 by (stac (real_minus_zero_less_iff RS sym) 1);
  1037 by (simp_tac (simpset() addsimps [real_add_commute,
  1038 				  real_less_sum_gt_0_iff]) 1);
  1039 qed "real_less_eq_diff";
  1040 
  1041 
  1042 (*** Subtraction laws ***)
  1043 
  1044 Goal "x + (y - z) = (x + y) - (z::real)";
  1045 by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
  1046 qed "real_add_diff_eq";
  1047 
  1048 Goal "(x - y) + z = (x + z) - (y::real)";
  1049 by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
  1050 qed "real_diff_add_eq";
  1051 
  1052 Goal "(x - y) - z = x - (y + (z::real))";
  1053 by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
  1054 qed "real_diff_diff_eq";
  1055 
  1056 Goal "x - (y - z) = (x + z) - (y::real)";
  1057 by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
  1058 qed "real_diff_diff_eq2";
  1059 
  1060 Goal "(x-y < z) = (x < z + (y::real))";
  1061 by (stac real_less_eq_diff 1);
  1062 by (res_inst_tac [("y1", "z")] (real_less_eq_diff RS ssubst) 1);
  1063 by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
  1064 qed "real_diff_less_eq";
  1065 
  1066 Goal "(x < z-y) = (x + (y::real) < z)";
  1067 by (stac real_less_eq_diff 1);
  1068 by (res_inst_tac [("y1", "z-y")] (real_less_eq_diff RS ssubst) 1);
  1069 by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
  1070 qed "real_less_diff_eq";
  1071 
  1072 Goalw [real_le_def] "(x-y <= z) = (x <= z + (y::real))";
  1073 by (simp_tac (simpset() addsimps [real_less_diff_eq]) 1);
  1074 qed "real_diff_le_eq";
  1075 
  1076 Goalw [real_le_def] "(x <= z-y) = (x + (y::real) <= z)";
  1077 by (simp_tac (simpset() addsimps [real_diff_less_eq]) 1);
  1078 qed "real_le_diff_eq";
  1079 
  1080 Goalw [real_diff_def] "(x-y = z) = (x = z + (y::real))";
  1081 by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
  1082 qed "real_diff_eq_eq";
  1083 
  1084 Goalw [real_diff_def] "(x = z-y) = (x + (y::real) = z)";
  1085 by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
  1086 qed "real_eq_diff_eq";
  1087 
  1088 (*This list of rewrites simplifies (in)equalities by bringing subtractions
  1089   to the top and then moving negative terms to the other side.  
  1090   Use with real_add_ac*)
  1091 val real_compare_rls = 
  1092   [symmetric real_diff_def,
  1093    real_add_diff_eq, real_diff_add_eq, real_diff_diff_eq, real_diff_diff_eq2, 
  1094    real_diff_less_eq, real_less_diff_eq, real_diff_le_eq, real_le_diff_eq, 
  1095    real_diff_eq_eq, real_eq_diff_eq];
  1096 
  1097 
  1098 (** For the cancellation simproc.
  1099     The idea is to cancel like terms on opposite sides by subtraction **)
  1100 
  1101 Goal "(x::real) - y = x' - y' ==> (x<y) = (x'<y')";
  1102 by (stac real_less_eq_diff 1);
  1103 by (res_inst_tac [("y1", "y")] (real_less_eq_diff RS ssubst) 1);
  1104 by (Asm_simp_tac 1);
  1105 qed "real_less_eqI";
  1106 
  1107 Goal "(x::real) - y = x' - y' ==> (y<=x) = (y'<=x')";
  1108 by (dtac real_less_eqI 1);
  1109 by (asm_simp_tac (simpset() addsimps [real_le_def]) 1);
  1110 qed "real_le_eqI";
  1111 
  1112 Goal "(x::real) - y = x' - y' ==> (x=y) = (x'=y')";
  1113 by Safe_tac;
  1114 by (ALLGOALS
  1115     (asm_full_simp_tac
  1116      (simpset() addsimps [real_eq_diff_eq, real_diff_eq_eq])));
  1117 qed "real_eq_eqI";