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