| 7334 |      1 | (*  Title:       HOL/Real/Real.ML
 | 
|  |      2 |     ID:          $Id$
 | 
|  |      3 |     Author:      Lawrence C. Paulson
 | 
|  |      4 |                  Jacques D. Fleuriot
 | 
|  |      5 |     Copyright:   1998  University of Cambridge
 | 
|  |      6 |     Description: Type "real" is a linear order
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | Goal "(0r < x) = (? y. x = real_of_preal y)";
 | 
|  |     10 | by (auto_tac (claset(), simpset() addsimps [real_of_preal_zero_less]));
 | 
|  |     11 | by (cut_inst_tac [("x","x")] real_of_preal_trichotomy 1);
 | 
|  |     12 | by (blast_tac (claset() addSEs [real_less_irrefl,
 | 
|  |     13 | 				real_of_preal_not_minus_gt_zero RS notE]) 1);
 | 
|  |     14 | qed "real_gt_zero_preal_Ex";
 | 
|  |     15 | 
 | 
|  |     16 | Goal "real_of_preal z < x ==> ? y. x = real_of_preal y";
 | 
|  |     17 | by (blast_tac (claset() addSDs [real_of_preal_zero_less RS real_less_trans]
 | 
|  |     18 |                addIs [real_gt_zero_preal_Ex RS iffD1]) 1);
 | 
|  |     19 | qed "real_gt_preal_preal_Ex";
 | 
|  |     20 | 
 | 
|  |     21 | Goal "real_of_preal z <= x ==> ? y. x = real_of_preal y";
 | 
|  |     22 | by (blast_tac (claset() addDs [real_le_imp_less_or_eq,
 | 
|  |     23 | 			       real_gt_preal_preal_Ex]) 1);
 | 
|  |     24 | qed "real_ge_preal_preal_Ex";
 | 
|  |     25 | 
 | 
|  |     26 | Goal "y <= 0r ==> !x. y < real_of_preal x";
 | 
|  |     27 | by (auto_tac (claset() addEs [real_le_imp_less_or_eq RS disjE]
 | 
|  |     28 |                        addIs [real_of_preal_zero_less RSN(2,real_less_trans)],
 | 
|  |     29 |               simpset() addsimps [real_of_preal_zero_less]));
 | 
|  |     30 | qed "real_less_all_preal";
 | 
|  |     31 | 
 | 
|  |     32 | Goal "~ 0r < y ==> !x. y < real_of_preal x";
 | 
|  |     33 | by (blast_tac (claset() addSIs [real_less_all_preal,real_leI]) 1);
 | 
|  |     34 | qed "real_less_all_real2";
 | 
|  |     35 | 
 | 
|  |     36 | Goal "((x::real) < y) = (-y < -x)";
 | 
|  |     37 | by (rtac (real_less_sum_gt_0_iff RS subst) 1);
 | 
|  |     38 | by (res_inst_tac [("W1","x")] (real_less_sum_gt_0_iff RS subst) 1);
 | 
|  |     39 | by (simp_tac (simpset() addsimps [real_add_commute]) 1);
 | 
|  |     40 | qed "real_less_swap_iff";
 | 
|  |     41 | 
 | 
|  |     42 | Goal "[| R + L = S; 0r < L |] ==> R < S";
 | 
|  |     43 | by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
 | 
|  |     44 | by (auto_tac (claset(), simpset() addsimps real_add_ac));
 | 
|  |     45 | qed "real_lemma_add_positive_imp_less";
 | 
|  |     46 | 
 | 
|  |     47 | Goal "? T. 0r < T & R + T = S ==> R < S";
 | 
|  |     48 | by (blast_tac (claset() addIs [real_lemma_add_positive_imp_less]) 1);
 | 
|  |     49 | qed "real_ex_add_positive_left_less";
 | 
|  |     50 | 
 | 
|  |     51 | (*Alternative definition for real_less.  NOT for rewriting*)
 | 
|  |     52 | Goal "(R < S) = (? T. 0r < T & R + T = S)";
 | 
|  |     53 | by (blast_tac (claset() addSIs [real_less_add_positive_left_Ex,
 | 
|  |     54 | 				real_ex_add_positive_left_less]) 1);
 | 
|  |     55 | qed "real_less_iff_add";
 | 
|  |     56 | 
 | 
|  |     57 | Goal "(0r < x) = (-x < x)";
 | 
|  |     58 | by Safe_tac;
 | 
|  |     59 | by (rtac ccontr 2 THEN forward_tac 
 | 
|  |     60 |     [real_leI RS real_le_imp_less_or_eq] 2);
 | 
|  |     61 | by (Step_tac 2);
 | 
|  |     62 | by (dtac (real_minus_zero_less_iff RS iffD2) 2);
 | 
|  |     63 | by (blast_tac (claset() addIs [real_less_trans]) 2);
 | 
|  |     64 | by (auto_tac (claset(),
 | 
|  |     65 | 	      simpset() addsimps 
 | 
|  |     66 |  	        [real_gt_zero_preal_Ex,real_of_preal_minus_less_self]));
 | 
|  |     67 | qed "real_gt_zero_iff";
 | 
|  |     68 | 
 | 
|  |     69 | Goal "(x < 0r) = (x < -x)";
 | 
|  |     70 | by (rtac (real_minus_zero_less_iff RS subst) 1);
 | 
|  |     71 | by (stac real_gt_zero_iff 1);
 | 
|  |     72 | by (Full_simp_tac 1);
 | 
|  |     73 | qed "real_lt_zero_iff";
 | 
|  |     74 | 
 | 
|  |     75 | Goalw [real_le_def] "(0r <= x) = (-x <= x)";
 | 
|  |     76 | by (auto_tac (claset(), simpset() addsimps [real_lt_zero_iff RS sym]));
 | 
|  |     77 | qed "real_ge_zero_iff";
 | 
|  |     78 | 
 | 
|  |     79 | Goalw [real_le_def] "(x <= 0r) = (x <= -x)";
 | 
|  |     80 | by (auto_tac (claset(), simpset() addsimps [real_gt_zero_iff RS sym]));
 | 
|  |     81 | qed "real_le_zero_iff";
 | 
|  |     82 | 
 | 
|  |     83 | Goal "(real_of_preal m1 <= real_of_preal m2) = (m1 <= m2)";
 | 
|  |     84 | by (auto_tac (claset() addSIs [preal_leI],
 | 
|  |     85 |     simpset() addsimps [real_less_le_iff RS sym]));
 | 
|  |     86 | by (dtac preal_le_less_trans 1 THEN assume_tac 1);
 | 
|  |     87 | by (etac preal_less_irrefl 1);
 | 
|  |     88 | qed "real_of_preal_le_iff";
 | 
|  |     89 | 
 | 
|  |     90 | Goal "[| 0r < x; 0r < y |] ==> 0r < x * y";
 | 
|  |     91 | by (auto_tac (claset(), simpset() addsimps [real_gt_zero_preal_Ex]));  
 | 
|  |     92 | by (res_inst_tac [("x","y*ya")] exI 1);
 | 
|  |     93 | by (full_simp_tac (simpset() addsimps [real_of_preal_mult]) 1);
 | 
|  |     94 | qed "real_mult_order";
 | 
|  |     95 | 
 | 
|  |     96 | Goal "[| x < 0r; y < 0r |] ==> 0r < x * y";
 | 
|  |     97 | by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1));
 | 
|  |     98 | by (dtac real_mult_order 1 THEN assume_tac 1);
 | 
|  |     99 | by (Asm_full_simp_tac 1);
 | 
|  |    100 | qed "real_mult_less_zero1";
 | 
|  |    101 | 
 | 
|  |    102 | Goal "[| 0r <= x; 0r <= y |] ==> 0r <= x * y";
 | 
|  |    103 | by (REPEAT(dtac real_le_imp_less_or_eq 1));
 | 
|  |    104 | by (auto_tac (claset() addIs [real_mult_order, real_less_imp_le],
 | 
|  |    105 | 	      simpset()));
 | 
|  |    106 | qed "real_le_mult_order";
 | 
|  |    107 | 
 | 
|  |    108 | Goal "[| 0r < x; 0r <= y |] ==> 0r <= x * y";
 | 
|  |    109 | by (dtac real_le_imp_less_or_eq 1);
 | 
|  |    110 | by (auto_tac (claset() addIs [real_mult_order,
 | 
|  |    111 | 			      real_less_imp_le],simpset()));
 | 
|  |    112 | qed "real_less_le_mult_order";
 | 
|  |    113 | 
 | 
|  |    114 | Goal "[| x <= 0r; y <= 0r |] ==> 0r <= x * y";
 | 
|  |    115 | by (rtac real_less_or_eq_imp_le 1);
 | 
|  |    116 | by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1);
 | 
|  |    117 | by Auto_tac;
 | 
|  |    118 | by (dtac real_le_imp_less_or_eq 1);
 | 
|  |    119 | by (auto_tac (claset() addDs [real_mult_less_zero1],simpset()));
 | 
|  |    120 | qed "real_mult_le_zero1";
 | 
|  |    121 | 
 | 
|  |    122 | Goal "[| 0r <= x; y < 0r |] ==> x * y <= 0r";
 | 
|  |    123 | by (rtac real_less_or_eq_imp_le 1);
 | 
|  |    124 | by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1);
 | 
|  |    125 | by Auto_tac;
 | 
|  |    126 | by (dtac (real_minus_zero_less_iff RS iffD2) 1);
 | 
|  |    127 | by (rtac (real_minus_zero_less_iff RS subst) 1);
 | 
|  |    128 | by (blast_tac (claset() addDs [real_mult_order] 
 | 
|  |    129 | 	                addIs [real_minus_mult_eq2 RS ssubst]) 1);
 | 
|  |    130 | qed "real_mult_le_zero";
 | 
|  |    131 | 
 | 
|  |    132 | Goal "[| 0r < x; y < 0r |] ==> x*y < 0r";
 | 
|  |    133 | by (dtac (real_minus_zero_less_iff RS iffD2) 1);
 | 
|  |    134 | by (dtac real_mult_order 1 THEN assume_tac 1);
 | 
|  |    135 | by (rtac (real_minus_zero_less_iff RS iffD1) 1);
 | 
|  |    136 | by (asm_full_simp_tac (simpset() addsimps [real_minus_mult_eq2]) 1);
 | 
|  |    137 | qed "real_mult_less_zero";
 | 
|  |    138 | 
 | 
|  |    139 | Goalw [real_one_def] "0r < 1r";
 | 
|  |    140 | by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2],
 | 
|  |    141 | 	      simpset() addsimps [real_of_preal_def]));
 | 
|  |    142 | qed "real_zero_less_one";
 | 
|  |    143 | 
 | 
|  |    144 | (*** Monotonicity results ***)
 | 
|  |    145 | 
 | 
|  |    146 | Goal "(v+z < w+z) = (v < (w::real))";
 | 
|  |    147 | by (Simp_tac 1);
 | 
|  |    148 | qed "real_add_right_cancel_less";
 | 
|  |    149 | 
 | 
|  |    150 | Goal "(z+v < z+w) = (v < (w::real))";
 | 
|  |    151 | by (Simp_tac 1);
 | 
|  |    152 | qed "real_add_left_cancel_less";
 | 
|  |    153 | 
 | 
|  |    154 | Addsimps [real_add_right_cancel_less, real_add_left_cancel_less];
 | 
|  |    155 | 
 | 
|  |    156 | Goal "(v+z <= w+z) = (v <= (w::real))";
 | 
|  |    157 | by (Simp_tac 1);
 | 
|  |    158 | qed "real_add_right_cancel_le";
 | 
|  |    159 | 
 | 
|  |    160 | Goal "(z+v <= z+w) = (v <= (w::real))";
 | 
|  |    161 | by (Simp_tac 1);
 | 
|  |    162 | qed "real_add_left_cancel_le";
 | 
|  |    163 | 
 | 
|  |    164 | Addsimps [real_add_right_cancel_le, real_add_left_cancel_le];
 | 
|  |    165 | 
 | 
|  |    166 | (*"v<=w ==> v+z <= w+z"*)
 | 
|  |    167 | bind_thm ("real_add_less_mono1", real_add_right_cancel_less RS iffD2);
 | 
|  |    168 | 
 | 
|  |    169 | (*"v<=w ==> v+z <= w+z"*)
 | 
|  |    170 | bind_thm ("real_add_le_mono1", real_add_right_cancel_le RS iffD2);
 | 
|  |    171 | 
 | 
|  |    172 | Goal "!!z z'::real. [| w'<w; z'<=z |] ==> w' + z' < w + z";
 | 
|  |    173 | by (etac (real_add_less_mono1 RS real_less_le_trans) 1);
 | 
|  |    174 | by (Simp_tac 1);
 | 
|  |    175 | qed "real_add_less_le_mono";
 | 
|  |    176 | 
 | 
|  |    177 | Goal "!!z z'::real. [| w'<=w; z'<z |] ==> w' + z' < w + z";
 | 
|  |    178 | by (etac (real_add_le_mono1 RS real_le_less_trans) 1);
 | 
|  |    179 | by (Simp_tac 1);
 | 
|  |    180 | qed "real_add_le_less_mono";
 | 
|  |    181 | 
 | 
|  |    182 | Goal "!!(A::real). A < B ==> C + A < C + B";
 | 
|  |    183 | by (Simp_tac 1);
 | 
|  |    184 | qed "real_add_less_mono2";
 | 
|  |    185 | 
 | 
|  |    186 | Goal "!!(A::real). A + C < B + C ==> A < B";
 | 
|  |    187 | by (Full_simp_tac 1);
 | 
|  |    188 | qed "real_less_add_right_cancel";
 | 
|  |    189 | 
 | 
|  |    190 | Goal "!!(A::real). C + A < C + B ==> A < B";
 | 
|  |    191 | by (Full_simp_tac 1);
 | 
|  |    192 | qed "real_less_add_left_cancel";
 | 
|  |    193 | 
 | 
|  |    194 | Goal "!!(A::real). A + C <= B + C ==> A <= B";
 | 
|  |    195 | by (Full_simp_tac 1);
 | 
|  |    196 | qed "real_le_add_right_cancel";
 | 
|  |    197 | 
 | 
|  |    198 | Goal "!!(A::real). C + A <= C + B ==> A <= B";
 | 
|  |    199 | by (Full_simp_tac 1);
 | 
|  |    200 | qed "real_le_add_left_cancel";
 | 
|  |    201 | 
 | 
|  |    202 | Goal "[| 0r < x; 0r < y |] ==> 0r < x + y";
 | 
|  |    203 | by (etac real_less_trans 1);
 | 
|  |    204 | by (dtac real_add_less_mono2 1);
 | 
|  |    205 | by (Full_simp_tac 1);
 | 
|  |    206 | qed "real_add_order";
 | 
|  |    207 | 
 | 
|  |    208 | Goal "[| 0r <= x; 0r <= y |] ==> 0r <= x + y";
 | 
|  |    209 | by (REPEAT(dtac real_le_imp_less_or_eq 1));
 | 
|  |    210 | by (auto_tac (claset() addIs [real_add_order, real_less_imp_le],
 | 
|  |    211 | 	      simpset()));
 | 
|  |    212 | qed "real_le_add_order";
 | 
|  |    213 | 
 | 
|  |    214 | Goal "[| R1 < S1; R2 < S2 |] ==> R1 + R2 < S1 + (S2::real)";
 | 
|  |    215 | by (dtac real_add_less_mono1 1);
 | 
|  |    216 | by (etac real_less_trans 1);
 | 
|  |    217 | by (etac real_add_less_mono2 1);
 | 
|  |    218 | qed "real_add_less_mono";
 | 
|  |    219 | 
 | 
|  |    220 | Goal "!!(q1::real). q1 <= q2  ==> x + q1 <= x + q2";
 | 
|  |    221 | by (Simp_tac 1);
 | 
|  |    222 | qed "real_add_left_le_mono1";
 | 
|  |    223 | 
 | 
|  |    224 | Goal "[|i<=j;  k<=l |] ==> i + k <= j + (l::real)";
 | 
|  |    225 | by (dtac real_add_le_mono1 1);
 | 
|  |    226 | by (etac real_le_trans 1);
 | 
|  |    227 | by (Simp_tac 1);
 | 
|  |    228 | qed "real_add_le_mono";
 | 
|  |    229 | 
 | 
|  |    230 | Goal "EX (x::real). x < y";
 | 
|  |    231 | by (rtac (real_add_zero_right RS subst) 1);
 | 
|  |    232 | by (res_inst_tac [("x","y + (-1r)")] exI 1);
 | 
|  |    233 | by (auto_tac (claset() addSIs [real_add_less_mono2],
 | 
|  |    234 | 	  simpset() addsimps [real_minus_zero_less_iff2, real_zero_less_one]));
 | 
|  |    235 | qed "real_less_Ex";
 | 
|  |    236 | 
 | 
|  |    237 | Goal "0r < r ==>  u + (-r) < u";
 | 
|  |    238 | by (res_inst_tac [("C","r")] real_less_add_right_cancel 1);
 | 
|  |    239 | by (simp_tac (simpset() addsimps [real_add_assoc]) 1);
 | 
|  |    240 | qed "real_add_minus_positive_less_self";
 | 
|  |    241 | 
 | 
|  |    242 | Goal "((r::real) <= s) = (-s <= -r)";
 | 
|  |    243 | by (Step_tac 1);
 | 
|  |    244 | by (dres_inst_tac [("x","-s")] real_add_left_le_mono1 1);
 | 
|  |    245 | by (dres_inst_tac [("x","r")] real_add_left_le_mono1 2);
 | 
|  |    246 | by Auto_tac;
 | 
|  |    247 | by (dres_inst_tac [("z","-r")] real_add_le_mono1 1);
 | 
|  |    248 | by (dres_inst_tac [("z","s")] real_add_le_mono1 2);
 | 
|  |    249 | by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
 | 
|  |    250 | qed "real_le_minus_iff";
 | 
|  |    251 | Addsimps [real_le_minus_iff RS sym];
 | 
|  |    252 |           
 | 
|  |    253 | Goal "0r <= 1r";
 | 
|  |    254 | by (rtac (real_zero_less_one RS real_less_imp_le) 1);
 | 
|  |    255 | qed "real_zero_le_one";
 | 
|  |    256 | Addsimps [real_zero_le_one];
 | 
|  |    257 | 
 | 
|  |    258 | Goal "0r <= x*x";
 | 
|  |    259 | by (res_inst_tac [("R2.0","0r"),("R1.0","x")] real_linear_less2 1);
 | 
|  |    260 | by (auto_tac (claset() addIs [real_mult_order,
 | 
|  |    261 | 			      real_mult_less_zero1,real_less_imp_le],
 | 
|  |    262 | 	      simpset()));
 | 
|  |    263 | qed "real_le_square";
 | 
|  |    264 | Addsimps [real_le_square];
 | 
|  |    265 | 
 | 
|  |    266 | (*----------------------------------------------------------------------------
 | 
|  |    267 |              An embedding of the naturals in the reals
 | 
|  |    268 |  ----------------------------------------------------------------------------*)
 | 
|  |    269 | 
 | 
|  |    270 | Goalw [real_of_posnat_def] "real_of_posnat 0 = 1r";
 | 
|  |    271 | by (full_simp_tac (simpset() addsimps [pnat_one_iff RS sym,real_of_preal_def]) 1);
 | 
|  |    272 | by (fold_tac [real_one_def]);
 | 
|  |    273 | by (rtac refl 1);
 | 
|  |    274 | qed "real_of_posnat_one";
 | 
|  |    275 | 
 | 
|  |    276 | Goalw [real_of_posnat_def] "real_of_posnat 1 = 1r + 1r";
 | 
|  |    277 | by (full_simp_tac (simpset() addsimps [real_of_preal_def,real_one_def,
 | 
|  |    278 |     pnat_two_eq,real_add,prat_of_pnat_add RS sym,preal_of_prat_add RS sym
 | 
|  |    279 |     ] @ pnat_add_ac) 1);
 | 
|  |    280 | qed "real_of_posnat_two";
 | 
|  |    281 | 
 | 
|  |    282 | Goalw [real_of_posnat_def]
 | 
|  |    283 |     "real_of_posnat n1 + real_of_posnat n2 = real_of_posnat (n1 + n2) + 1r";
 | 
|  |    284 | by (full_simp_tac (simpset() addsimps [real_of_posnat_one RS sym,
 | 
|  |    285 |     real_of_posnat_def,real_of_preal_add RS sym,preal_of_prat_add RS sym,
 | 
|  |    286 |     prat_of_pnat_add RS sym,pnat_of_nat_add]) 1);
 | 
|  |    287 | qed "real_of_posnat_add";
 | 
|  |    288 | 
 | 
|  |    289 | Goal "real_of_posnat (n + 1) = real_of_posnat n + 1r";
 | 
|  |    290 | by (res_inst_tac [("x1","1r")] (real_add_right_cancel RS iffD1) 1);
 | 
|  |    291 | by (rtac (real_of_posnat_add RS subst) 1);
 | 
|  |    292 | by (full_simp_tac (simpset() addsimps [real_of_posnat_two,real_add_assoc]) 1);
 | 
|  |    293 | qed "real_of_posnat_add_one";
 | 
|  |    294 | 
 | 
|  |    295 | Goal "real_of_posnat (Suc n) = real_of_posnat n + 1r";
 | 
|  |    296 | by (stac (real_of_posnat_add_one RS sym) 1);
 | 
|  |    297 | by (Simp_tac 1);
 | 
|  |    298 | qed "real_of_posnat_Suc";
 | 
|  |    299 | 
 | 
|  |    300 | Goal "inj(real_of_posnat)";
 | 
|  |    301 | by (rtac injI 1);
 | 
|  |    302 | by (rewtac real_of_posnat_def);
 | 
|  |    303 | by (dtac (inj_real_of_preal RS injD) 1);
 | 
|  |    304 | by (dtac (inj_preal_of_prat RS injD) 1);
 | 
|  |    305 | by (dtac (inj_prat_of_pnat RS injD) 1);
 | 
|  |    306 | by (etac (inj_pnat_of_nat RS injD) 1);
 | 
|  |    307 | qed "inj_real_of_posnat";
 | 
|  |    308 | 
 | 
|  |    309 | Goalw [real_of_posnat_def] "0r < real_of_posnat n";
 | 
|  |    310 | by (rtac (real_gt_zero_preal_Ex RS iffD2) 1);
 | 
|  |    311 | by (Blast_tac 1);
 | 
|  |    312 | qed "real_of_posnat_less_zero";
 | 
|  |    313 | 
 | 
|  |    314 | Goal "real_of_posnat n ~= 0r";
 | 
|  |    315 | by (rtac (real_of_posnat_less_zero RS real_not_refl2 RS not_sym) 1);
 | 
|  |    316 | qed "real_of_posnat_not_eq_zero";
 | 
|  |    317 | Addsimps[real_of_posnat_not_eq_zero];
 | 
|  |    318 | 
 | 
|  |    319 | Goal "1r <= real_of_posnat n";
 | 
|  |    320 | by (simp_tac (simpset() addsimps [real_of_posnat_one RS sym]) 1);
 | 
|  |    321 | by (induct_tac "n" 1);
 | 
|  |    322 | by (auto_tac (claset(),
 | 
|  |    323 | 	      simpset () addsimps [real_of_posnat_Suc,real_of_posnat_one,
 | 
|  |    324 | 			   real_of_posnat_less_zero, real_less_imp_le]));
 | 
|  |    325 | qed "real_of_posnat_less_one";
 | 
|  |    326 | Addsimps [real_of_posnat_less_one];
 | 
|  |    327 | 
 | 
|  |    328 | Goal "rinv(real_of_posnat n) ~= 0r";
 | 
|  |    329 | by (rtac ((real_of_posnat_less_zero RS 
 | 
|  |    330 |     real_not_refl2 RS not_sym) RS rinv_not_zero) 1);
 | 
|  |    331 | qed "real_of_posnat_rinv_not_zero";
 | 
|  |    332 | Addsimps [real_of_posnat_rinv_not_zero];
 | 
|  |    333 | 
 | 
|  |    334 | Goal "rinv(real_of_posnat x) = rinv(real_of_posnat y) ==> x = y";
 | 
|  |    335 | by (rtac (inj_real_of_posnat RS injD) 1);
 | 
|  |    336 | by (res_inst_tac [("n2","x")] 
 | 
|  |    337 |     (real_of_posnat_rinv_not_zero RS real_mult_left_cancel RS iffD1) 1);
 | 
|  |    338 | by (full_simp_tac (simpset() addsimps [(real_of_posnat_less_zero RS 
 | 
|  |    339 |     real_not_refl2 RS not_sym) RS real_mult_inv_left]) 1);
 | 
|  |    340 | by (asm_full_simp_tac (simpset() addsimps [(real_of_posnat_less_zero RS 
 | 
|  |    341 |     real_not_refl2 RS not_sym)]) 1);
 | 
|  |    342 | qed "real_of_posnat_rinv_inj";
 | 
|  |    343 | 
 | 
|  |    344 | Goal "0r < x ==> 0r < rinv x";
 | 
|  |    345 | by (EVERY1[rtac ccontr, dtac real_leI]);
 | 
|  |    346 | by (forward_tac [real_minus_zero_less_iff2 RS iffD2] 1);
 | 
|  |    347 | by (forward_tac [real_not_refl2 RS not_sym] 1);
 | 
|  |    348 | by (dtac (real_not_refl2 RS not_sym RS rinv_not_zero) 1);
 | 
|  |    349 | by (EVERY1[dtac real_le_imp_less_or_eq, Step_tac]); 
 | 
|  |    350 | by (dtac real_mult_less_zero1 1 THEN assume_tac 1);
 | 
|  |    351 | by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym],
 | 
|  |    352 | 	      simpset() addsimps [real_minus_mult_eq1 RS sym]));
 | 
|  |    353 | qed "real_rinv_gt_zero";
 | 
|  |    354 | 
 | 
|  |    355 | Goal "x < 0r ==> rinv x < 0r";
 | 
| 7499 |    356 | by (ftac real_not_refl2 1);
 | 
| 7334 |    357 | by (dtac (real_minus_zero_less_iff RS iffD2) 1);
 | 
|  |    358 | by (rtac (real_minus_zero_less_iff RS iffD1) 1);
 | 
|  |    359 | by (dtac (real_minus_rinv RS sym) 1);
 | 
|  |    360 | by (auto_tac (claset() addIs [real_rinv_gt_zero], simpset()));
 | 
|  |    361 | qed "real_rinv_less_zero";
 | 
|  |    362 | 
 | 
|  |    363 | Goal "0r < rinv(real_of_posnat n)";
 | 
|  |    364 | by (rtac (real_of_posnat_less_zero RS real_rinv_gt_zero) 1);
 | 
|  |    365 | qed "real_of_posnat_rinv_gt_zero";
 | 
|  |    366 | Addsimps [real_of_posnat_rinv_gt_zero];
 | 
|  |    367 | 
 | 
|  |    368 | Goal "x+x = x*(1r+1r)";
 | 
|  |    369 | by (simp_tac (simpset() addsimps 
 | 
|  |    370 |               [real_add_mult_distrib2]) 1);
 | 
|  |    371 | qed "real_add_self";
 | 
|  |    372 | 
 | 
|  |    373 | Goal "x < x + 1r";
 | 
|  |    374 | by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
 | 
|  |    375 | by (full_simp_tac (simpset() addsimps [real_zero_less_one,
 | 
|  |    376 | 				real_add_assoc, real_add_left_commute]) 1);
 | 
|  |    377 | qed "real_self_less_add_one";
 | 
|  |    378 | 
 | 
|  |    379 | Goal "1r < 1r + 1r";
 | 
|  |    380 | by (rtac real_self_less_add_one 1);
 | 
|  |    381 | qed "real_one_less_two";
 | 
|  |    382 | 
 | 
|  |    383 | Goal "0r < 1r + 1r";
 | 
|  |    384 | by (rtac ([real_zero_less_one,
 | 
|  |    385 | 	   real_one_less_two] MRS real_less_trans) 1);
 | 
|  |    386 | qed "real_zero_less_two";
 | 
|  |    387 | 
 | 
|  |    388 | Goal "1r + 1r ~= 0r";
 | 
|  |    389 | by (rtac (real_zero_less_two RS real_not_refl2 RS not_sym) 1);
 | 
|  |    390 | qed "real_two_not_zero";
 | 
|  |    391 | 
 | 
|  |    392 | Addsimps [real_two_not_zero];
 | 
|  |    393 | 
 | 
|  |    394 | Goal "x*rinv(1r + 1r) + x*rinv(1r + 1r) = x";
 | 
|  |    395 | by (stac real_add_self 1);
 | 
|  |    396 | by (full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
 | 
|  |    397 | qed "real_sum_of_halves";
 | 
|  |    398 | 
 | 
|  |    399 | Goal "[| 0r<z; x<y |] ==> x*z<y*z";       
 | 
|  |    400 | by (rotate_tac 1 1);
 | 
|  |    401 | by (dtac real_less_sum_gt_zero 1);
 | 
|  |    402 | by (rtac real_sum_gt_zero_less 1);
 | 
|  |    403 | by (dtac real_mult_order 1 THEN assume_tac 1);
 | 
|  |    404 | by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2,
 | 
|  |    405 |     real_minus_mult_eq2 RS sym, real_mult_commute ]) 1);
 | 
|  |    406 | qed "real_mult_less_mono1";
 | 
|  |    407 | 
 | 
|  |    408 | Goal "[| 0r<z; x<y |] ==> z*x<z*y";       
 | 
|  |    409 | by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_less_mono1]) 1);
 | 
|  |    410 | qed "real_mult_less_mono2";
 | 
|  |    411 | 
 | 
|  |    412 | Goal "[| 0r<z; x*z<y*z |] ==> x<y";
 | 
|  |    413 | by (forw_inst_tac [("x","x*z")] (real_rinv_gt_zero 
 | 
|  |    414 |                        RS real_mult_less_mono1) 1);
 | 
|  |    415 | by (auto_tac (claset(),
 | 
|  |    416 | 	      simpset() addsimps 
 | 
|  |    417 |      [real_mult_assoc,real_not_refl2 RS not_sym]));
 | 
|  |    418 | qed "real_mult_less_cancel1";
 | 
|  |    419 | 
 | 
|  |    420 | Goal "[| 0r<z; z*x<z*y |] ==> x<y";
 | 
|  |    421 | by (etac real_mult_less_cancel1 1);
 | 
|  |    422 | by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1);
 | 
|  |    423 | qed "real_mult_less_cancel2";
 | 
|  |    424 | 
 | 
|  |    425 | Goal "0r < z ==> (x*z < y*z) = (x < y)";
 | 
|  |    426 | by (blast_tac (claset() addIs [real_mult_less_mono1,
 | 
|  |    427 |     real_mult_less_cancel1]) 1);
 | 
|  |    428 | qed "real_mult_less_iff1";
 | 
|  |    429 | 
 | 
|  |    430 | Goal "0r < z ==> (z*x < z*y) = (x < y)";
 | 
|  |    431 | by (blast_tac (claset() addIs [real_mult_less_mono2,
 | 
|  |    432 |     real_mult_less_cancel2]) 1);
 | 
|  |    433 | qed "real_mult_less_iff2";
 | 
|  |    434 | 
 | 
|  |    435 | Addsimps [real_mult_less_iff1,real_mult_less_iff2];
 | 
|  |    436 | 
 | 
|  |    437 | Goal "[| 0r<=z; x<y |] ==> x*z<=y*z";
 | 
|  |    438 | by (EVERY1 [rtac real_less_or_eq_imp_le, dtac real_le_imp_less_or_eq]);
 | 
|  |    439 | by (auto_tac (claset() addIs [real_mult_less_mono1],simpset()));
 | 
|  |    440 | qed "real_mult_le_less_mono1";
 | 
|  |    441 | 
 | 
|  |    442 | Goal "[| 0r<=z; x<y |] ==> z*x<=z*y";
 | 
|  |    443 | by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_le_less_mono1]) 1);
 | 
|  |    444 | qed "real_mult_le_less_mono2";
 | 
|  |    445 | 
 | 
|  |    446 | Goal "[| 0r<=z; x<=y |] ==> z*x<=z*y";
 | 
|  |    447 | by (dres_inst_tac [("x","x")] real_le_imp_less_or_eq 1);
 | 
|  |    448 | by (auto_tac (claset() addIs [real_mult_le_less_mono2], simpset()));
 | 
|  |    449 | qed "real_mult_le_le_mono1";
 | 
|  |    450 | 
 | 
|  |    451 | Goal "[| 0r < r1; r1 <r2; 0r < x; x < y|] ==> r1 * x < r2 * y";
 | 
|  |    452 | by (dres_inst_tac [("x","x")] real_mult_less_mono2 1);
 | 
|  |    453 | by (dres_inst_tac [("R1.0","0r")] real_less_trans 2);
 | 
|  |    454 | by (dres_inst_tac [("x","r1")] real_mult_less_mono1 3);
 | 
|  |    455 | by Auto_tac;
 | 
|  |    456 | by (blast_tac (claset() addIs [real_less_trans]) 1);
 | 
|  |    457 | qed "real_mult_less_mono";
 | 
|  |    458 | 
 | 
|  |    459 | Goal "[| 0r < r1; r1 <r2; 0r < y|] ==> 0r < r2 * y";
 | 
|  |    460 | by (dres_inst_tac [("R1.0","0r")] real_less_trans 1);
 | 
|  |    461 | by (assume_tac 1);
 | 
|  |    462 | by (blast_tac (claset() addIs [real_mult_order]) 1);
 | 
|  |    463 | qed "real_mult_order_trans";
 | 
|  |    464 | 
 | 
|  |    465 | Goal "[| 0r < r1; r1 <r2; 0r <= x; x < y|] ==> r1 * x < r2 * y";
 | 
|  |    466 | by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] 
 | 
|  |    467 | 	               addIs [real_mult_less_mono,real_mult_order_trans],
 | 
|  |    468 |               simpset()));
 | 
|  |    469 | qed "real_mult_less_mono3";
 | 
|  |    470 | 
 | 
|  |    471 | Goal "[| 0r <= r1; r1 <r2; 0r <= x; x < y|] ==> r1 * x < r2 * y";
 | 
|  |    472 | by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] 
 | 
|  |    473 | 	               addIs [real_mult_less_mono,real_mult_order_trans,
 | 
|  |    474 | 			      real_mult_order],
 | 
|  |    475 | 	      simpset()));
 | 
|  |    476 | by (dres_inst_tac [("R2.0","x")] real_less_trans 1);
 | 
|  |    477 | by (assume_tac 1);
 | 
|  |    478 | by (blast_tac (claset() addIs [real_mult_order]) 1);
 | 
|  |    479 | qed "real_mult_less_mono4";
 | 
|  |    480 | 
 | 
|  |    481 | Goal "[| 0r < r1; r1 <= r2; 0r <= x; x <= y |] ==> r1 * x <= r2 * y";
 | 
|  |    482 | by (rtac real_less_or_eq_imp_le 1);
 | 
|  |    483 | by (REPEAT(dtac real_le_imp_less_or_eq 1));
 | 
|  |    484 | by (auto_tac (claset() addIs [real_mult_less_mono,
 | 
|  |    485 | 			      real_mult_order_trans,real_mult_order],
 | 
|  |    486 | 	      simpset()));
 | 
|  |    487 | qed "real_mult_le_mono";
 | 
|  |    488 | 
 | 
|  |    489 | Goal "[| 0r < r1; r1 < r2; 0r <= x; x <= y |] ==> r1 * x <= r2 * y";
 | 
|  |    490 | by (rtac real_less_or_eq_imp_le 1);
 | 
|  |    491 | by (REPEAT(dtac real_le_imp_less_or_eq 1));
 | 
|  |    492 | by (auto_tac (claset() addIs [real_mult_less_mono, real_mult_order_trans,
 | 
|  |    493 | 			      real_mult_order],
 | 
|  |    494 | 	      simpset()));
 | 
|  |    495 | qed "real_mult_le_mono2";
 | 
|  |    496 | 
 | 
|  |    497 | Goal "[| 0r <= r1; r1 < r2; 0r <= x; x <= y |] ==> r1 * x <= r2 * y";
 | 
|  |    498 | by (dtac real_le_imp_less_or_eq 1);
 | 
|  |    499 | by (auto_tac (claset() addIs [real_mult_le_mono2],simpset()));
 | 
|  |    500 | by (dtac real_le_trans 1 THEN assume_tac 1);
 | 
|  |    501 | by (auto_tac (claset() addIs [real_less_le_mult_order], simpset()));
 | 
|  |    502 | qed "real_mult_le_mono3";
 | 
|  |    503 | 
 | 
|  |    504 | Goal "[| 0r <= r1; r1 <= r2; 0r <= x; x <= y |] ==> r1 * x <= r2 * y";
 | 
|  |    505 | by (dres_inst_tac [("x","r1")] real_le_imp_less_or_eq 1);
 | 
|  |    506 | by (auto_tac (claset() addIs [real_mult_le_mono3, real_mult_le_le_mono1],
 | 
|  |    507 | 	      simpset()));
 | 
|  |    508 | qed "real_mult_le_mono4";
 | 
|  |    509 | 
 | 
|  |    510 | Goal "1r <= x ==> 0r < x";
 | 
|  |    511 | by (rtac ccontr 1 THEN dtac real_leI 1);
 | 
|  |    512 | by (dtac real_le_trans 1 THEN assume_tac 1);
 | 
|  |    513 | by (auto_tac (claset() addDs [real_zero_less_one RSN (2,real_le_less_trans)],
 | 
|  |    514 | 	      simpset() addsimps [real_less_not_refl]));
 | 
|  |    515 | qed "real_gt_zero";
 | 
|  |    516 | 
 | 
|  |    517 | Goal "[| 1r < r; 1r <= x |]  ==> x <= r * x";
 | 
|  |    518 | by (dtac (real_gt_zero RS real_less_imp_le) 1);
 | 
|  |    519 | by (auto_tac (claset() addSDs [real_mult_le_less_mono1],
 | 
|  |    520 |     simpset()));
 | 
|  |    521 | qed "real_mult_self_le";
 | 
|  |    522 | 
 | 
|  |    523 | Goal "[| 1r <= r; 1r <= x |]  ==> x <= r * x";
 | 
|  |    524 | by (dtac real_le_imp_less_or_eq 1);
 | 
|  |    525 | by (auto_tac (claset() addIs [real_mult_self_le],
 | 
|  |    526 | 	      simpset() addsimps [real_le_refl]));
 | 
|  |    527 | qed "real_mult_self_le2";
 | 
|  |    528 | 
 | 
|  |    529 | Goal "x < y ==> x < (x + y)*rinv(1r + 1r)";
 | 
|  |    530 | by (dres_inst_tac [("C","x")] real_add_less_mono2 1);
 | 
|  |    531 | by (dtac (real_add_self RS subst) 1);
 | 
|  |    532 | by (dtac (real_zero_less_two RS real_rinv_gt_zero RS 
 | 
|  |    533 |           real_mult_less_mono1) 1);
 | 
|  |    534 | by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
 | 
|  |    535 | qed "real_less_half_sum";
 | 
|  |    536 | 
 | 
|  |    537 | Goal "x < y ==> (x + y)*rinv(1r + 1r) < y";
 | 
|  |    538 | by (dtac real_add_less_mono1 1);
 | 
|  |    539 | by (dtac (real_add_self RS subst) 1);
 | 
|  |    540 | by (dtac (real_zero_less_two RS real_rinv_gt_zero RS 
 | 
|  |    541 |           real_mult_less_mono1) 1);
 | 
|  |    542 | by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
 | 
|  |    543 | qed "real_gt_half_sum";
 | 
|  |    544 | 
 | 
|  |    545 | Goal "x < y ==> EX r::real. x < r & r < y";
 | 
|  |    546 | by (blast_tac (claset() addSIs [real_less_half_sum,
 | 
|  |    547 | 				real_gt_half_sum]) 1);
 | 
|  |    548 | qed "real_dense";
 | 
|  |    549 | 
 | 
|  |    550 | Goal "(EX n. rinv(real_of_posnat n) < r) = (EX n. 1r < r * real_of_posnat n)";
 | 
|  |    551 | by (Step_tac 1);
 | 
|  |    552 | by (dres_inst_tac [("n1","n")] (real_of_posnat_less_zero 
 | 
|  |    553 | 				RS real_mult_less_mono1) 1);
 | 
|  |    554 | by (dres_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
 | 
|  |    555 | 				real_rinv_gt_zero RS real_mult_less_mono1) 2);
 | 
|  |    556 | by (auto_tac (claset(),
 | 
|  |    557 | 	      simpset() addsimps [(real_of_posnat_less_zero RS 
 | 
|  |    558 | 				   real_not_refl2 RS not_sym),
 | 
|  |    559 | 				  real_mult_assoc]));
 | 
|  |    560 | qed "real_of_posnat_rinv_Ex_iff";
 | 
|  |    561 | 
 | 
|  |    562 | Goal "(rinv(real_of_posnat n) < r) = (1r < r * real_of_posnat n)";
 | 
|  |    563 | by (Step_tac 1);
 | 
|  |    564 | by (dres_inst_tac [("n1","n")] (real_of_posnat_less_zero 
 | 
|  |    565 |                        RS real_mult_less_mono1) 1);
 | 
|  |    566 | by (dres_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
 | 
|  |    567 | 				real_rinv_gt_zero RS real_mult_less_mono1) 2);
 | 
|  |    568 | by (auto_tac (claset(), simpset() addsimps [real_mult_assoc]));
 | 
|  |    569 | qed "real_of_posnat_rinv_iff";
 | 
|  |    570 | 
 | 
|  |    571 | Goal "(rinv(real_of_posnat n) <= r) = (1r <= r * real_of_posnat n)";
 | 
|  |    572 | by (Step_tac 1);
 | 
|  |    573 | by (dres_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
 | 
|  |    574 |     real_less_imp_le RS real_mult_le_le_mono1) 1);
 | 
|  |    575 | by (dres_inst_tac [("n3","n")] (real_of_posnat_less_zero RS 
 | 
|  |    576 |         real_rinv_gt_zero RS real_less_imp_le RS 
 | 
|  |    577 |         real_mult_le_le_mono1) 2);
 | 
|  |    578 | by (auto_tac (claset(), simpset() addsimps real_mult_ac));
 | 
|  |    579 | qed "real_of_posnat_rinv_le_iff";
 | 
|  |    580 | 
 | 
|  |    581 | Goalw [real_of_posnat_def] "(real_of_posnat n < real_of_posnat m) = (n < m)";
 | 
|  |    582 | by Auto_tac;
 | 
|  |    583 | qed "real_of_posnat_less_iff";
 | 
|  |    584 | 
 | 
|  |    585 | Addsimps [real_of_posnat_less_iff];
 | 
|  |    586 | 
 | 
|  |    587 | Goal "0r < u  ==> (u < rinv (real_of_posnat n)) = (real_of_posnat n < rinv(u))";
 | 
|  |    588 | by (Step_tac 1);
 | 
|  |    589 | by (res_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
 | 
|  |    590 |     real_rinv_gt_zero RS real_mult_less_cancel1) 1);
 | 
|  |    591 | by (res_inst_tac [("x1","u")] ( real_rinv_gt_zero
 | 
|  |    592 |    RS real_mult_less_cancel1) 2);
 | 
|  |    593 | by (auto_tac (claset(),
 | 
|  |    594 | 	      simpset() addsimps [real_of_posnat_less_zero, 
 | 
|  |    595 |     real_not_refl2 RS not_sym]));
 | 
|  |    596 | by (res_inst_tac [("z","u")] real_mult_less_cancel2 1);
 | 
|  |    597 | by (res_inst_tac [("n1","n")] (real_of_posnat_less_zero RS 
 | 
|  |    598 |     real_mult_less_cancel2) 3);
 | 
|  |    599 | by (auto_tac (claset(),
 | 
|  |    600 | 	      simpset() addsimps [real_of_posnat_less_zero, 
 | 
|  |    601 |     real_not_refl2 RS not_sym,real_mult_assoc RS sym]));
 | 
|  |    602 | qed "real_of_posnat_less_rinv_iff";
 | 
|  |    603 | 
 | 
|  |    604 | Goal "0r < u ==> (u = rinv(real_of_posnat n)) = (real_of_posnat n = rinv u)";
 | 
|  |    605 | by (auto_tac (claset(),
 | 
|  |    606 | 	      simpset() addsimps [real_rinv_rinv,
 | 
|  |    607 |     real_of_posnat_less_zero,real_not_refl2 RS not_sym]));
 | 
|  |    608 | qed "real_of_posnat_rinv_eq_iff";
 | 
|  |    609 | 
 | 
|  |    610 | Goal "[| 0r < r; r < x |] ==> rinv x < rinv r";
 | 
| 7499 |    611 | by (ftac real_less_trans 1 THEN assume_tac 1);
 | 
|  |    612 | by (ftac real_rinv_gt_zero 1);
 | 
| 7334 |    613 | by (forw_inst_tac [("x","x")] real_rinv_gt_zero 1);
 | 
|  |    614 | by (forw_inst_tac [("x","r"),("z","rinv r")] real_mult_less_mono1 1);
 | 
|  |    615 | by (assume_tac 1);
 | 
|  |    616 | by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS 
 | 
|  |    617 | 					   not_sym RS real_mult_inv_right]) 1);
 | 
| 7499 |    618 | by (ftac real_rinv_gt_zero 1);
 | 
| 7334 |    619 | by (forw_inst_tac [("x","1r"),("z","rinv x")] real_mult_less_mono2 1);
 | 
|  |    620 | by (assume_tac 1);
 | 
|  |    621 | by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS 
 | 
|  |    622 |          not_sym RS real_mult_inv_left,real_mult_assoc RS sym]) 1);
 | 
|  |    623 | qed "real_rinv_less_swap";
 | 
|  |    624 | 
 | 
|  |    625 | Goal "[| 0r < r; 0r < x|] ==> (r < x) = (rinv x < rinv r)";
 | 
|  |    626 | by (auto_tac (claset() addIs [real_rinv_less_swap],simpset()));
 | 
|  |    627 | by (res_inst_tac [("t","r")] (real_rinv_rinv RS subst) 1);
 | 
|  |    628 | by (etac (real_not_refl2 RS not_sym) 1);
 | 
|  |    629 | by (res_inst_tac [("t","x")] (real_rinv_rinv RS subst) 1);
 | 
|  |    630 | by (etac (real_not_refl2 RS not_sym) 1);
 | 
|  |    631 | by (auto_tac (claset() addIs [real_rinv_less_swap],
 | 
|  |    632 | 	      simpset() addsimps [real_rinv_gt_zero]));
 | 
|  |    633 | qed "real_rinv_less_iff";
 | 
|  |    634 | 
 | 
|  |    635 | Goal "r < r + rinv(real_of_posnat n)";
 | 
|  |    636 | by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
 | 
|  |    637 | by (full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
 | 
|  |    638 | qed "real_add_rinv_real_of_posnat_less";
 | 
|  |    639 | Addsimps [real_add_rinv_real_of_posnat_less];
 | 
|  |    640 | 
 | 
|  |    641 | Goal "r <= r + rinv(real_of_posnat n)";
 | 
|  |    642 | by (rtac real_less_imp_le 1);
 | 
|  |    643 | by (Simp_tac 1);
 | 
|  |    644 | qed "real_add_rinv_real_of_posnat_le";
 | 
|  |    645 | Addsimps [real_add_rinv_real_of_posnat_le];
 | 
|  |    646 | 
 | 
|  |    647 | Goal "r + (-rinv(real_of_posnat n)) < r";
 | 
|  |    648 | by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
 | 
|  |    649 | by (full_simp_tac (simpset() addsimps [real_add_assoc RS sym,
 | 
|  |    650 | 				       real_minus_zero_less_iff2]) 1);
 | 
|  |    651 | qed "real_add_minus_rinv_real_of_posnat_less";
 | 
|  |    652 | Addsimps [real_add_minus_rinv_real_of_posnat_less];
 | 
|  |    653 | 
 | 
|  |    654 | Goal "r + (-rinv(real_of_posnat n)) <= r";
 | 
|  |    655 | by (rtac real_less_imp_le 1);
 | 
|  |    656 | by (Simp_tac 1);
 | 
|  |    657 | qed "real_add_minus_rinv_real_of_posnat_le";
 | 
|  |    658 | Addsimps [real_add_minus_rinv_real_of_posnat_le];
 | 
|  |    659 | 
 | 
|  |    660 | Goal "0r < r ==> r*(1r + (-rinv(real_of_posnat n))) < r";
 | 
|  |    661 | by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1);
 | 
|  |    662 | by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
 | 
|  |    663 | by (auto_tac (claset() addIs [real_mult_order],
 | 
|  |    664 | 	      simpset() addsimps [real_add_assoc RS sym,
 | 
|  |    665 | 				  real_minus_mult_eq2 RS sym,
 | 
|  |    666 | 				  real_minus_zero_less_iff2]));
 | 
|  |    667 | qed "real_mult_less_self";
 | 
|  |    668 | 
 | 
|  |    669 | Goal "0r <= 1r + (-rinv(real_of_posnat n))";
 | 
|  |    670 | by (res_inst_tac [("C","rinv(real_of_posnat n)")] real_le_add_right_cancel 1);
 | 
|  |    671 | by (simp_tac (simpset() addsimps [real_add_assoc,
 | 
|  |    672 | 				  real_of_posnat_rinv_le_iff]) 1);
 | 
|  |    673 | qed "real_add_one_minus_rinv_ge_zero";
 | 
|  |    674 | 
 | 
|  |    675 | Goal "0r < r ==> 0r <= r*(1r + (-rinv(real_of_posnat n)))";
 | 
|  |    676 | by (dtac (real_add_one_minus_rinv_ge_zero RS real_mult_le_less_mono1) 1);
 | 
|  |    677 | by Auto_tac;
 | 
|  |    678 | qed "real_mult_add_one_minus_ge_zero";
 | 
|  |    679 | 
 | 
|  |    680 | Goal "x*y = 0r ==> x = 0r | y = 0r";
 | 
|  |    681 | by (auto_tac (claset() addIs [ccontr] addDs [real_mult_not_zero],
 | 
|  |    682 | 	      simpset()));
 | 
|  |    683 | qed "real_mult_zero_disj";
 | 
|  |    684 |  
 | 
|  |    685 | Goal "x + x*y = x*(1r + y)";
 | 
|  |    686 | by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1);
 | 
|  |    687 | qed "real_add_mult_distrib_one";
 | 
|  |    688 | 
 | 
|  |    689 | Goal "[| x ~= 1r; y * x = y |] ==> y = 0r";
 | 
|  |    690 | by (dtac (sym RS (real_eq_minus_iff RS iffD1)) 1);
 | 
|  |    691 | by (dtac sym 1);
 | 
|  |    692 | by (asm_full_simp_tac (simpset() addsimps [real_minus_mult_eq2,
 | 
|  |    693 |     real_add_mult_distrib_one]) 1);
 | 
|  |    694 | by (dtac real_mult_zero_disj 1);
 | 
|  |    695 | by (auto_tac (claset(),
 | 
|  |    696 | 	      simpset() addsimps [real_eq_minus_iff2 RS sym]));
 | 
|  |    697 | qed "real_mult_eq_self_zero";
 | 
|  |    698 | Addsimps [real_mult_eq_self_zero];
 | 
|  |    699 | 
 | 
|  |    700 | Goal "[| x ~= 1r; y = y * x |] ==> y = 0r";
 | 
|  |    701 | by (dtac sym 1);
 | 
|  |    702 | by (Asm_full_simp_tac 1);
 | 
|  |    703 | qed "real_mult_eq_self_zero2";
 | 
|  |    704 | Addsimps [real_mult_eq_self_zero2];
 | 
|  |    705 | 
 | 
|  |    706 | Goal "[| 0r <= x*y; 0r < x |] ==> 0r <= y";
 | 
| 7499 |    707 | by (ftac real_rinv_gt_zero 1);
 | 
| 7334 |    708 | by (dres_inst_tac [("x","rinv x")] real_less_le_mult_order 1);
 | 
|  |    709 | by (dtac (real_not_refl2 RS not_sym RS real_mult_inv_left) 2);
 | 
|  |    710 | by (auto_tac (claset(),
 | 
|  |    711 | 	      simpset() addsimps [real_mult_assoc RS sym]));
 | 
|  |    712 | qed "real_mult_ge_zero_cancel";
 | 
|  |    713 | 
 | 
|  |    714 | Goal "[|x ~= 0r; y ~= 0r |] ==> rinv(x) + rinv(y) = (x + y)*rinv(x*y)";
 | 
|  |    715 | by (asm_full_simp_tac (simpset() addsimps 
 | 
|  |    716 | 		       [real_rinv_distrib,real_add_mult_distrib,
 | 
|  |    717 | 			real_mult_assoc RS sym]) 1);
 | 
|  |    718 | by (stac real_mult_assoc 1);
 | 
|  |    719 | by (rtac (real_mult_left_commute RS subst) 1);
 | 
|  |    720 | by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
 | 
|  |    721 | qed "real_rinv_add";
 | 
|  |    722 | 
 | 
|  |    723 | (*----------------------------------------------------------------------------
 | 
|  |    724 |      Another embedding of the naturals in the reals (see real_of_posnat)
 | 
|  |    725 |  ----------------------------------------------------------------------------*)
 | 
|  |    726 | Goalw [real_of_nat_def] "real_of_nat 0 = 0r";
 | 
|  |    727 | by (full_simp_tac (simpset() addsimps [real_of_posnat_one]) 1);
 | 
|  |    728 | qed "real_of_nat_zero";
 | 
|  |    729 | 
 | 
|  |    730 | Goalw [real_of_nat_def] "real_of_nat 1 = 1r";
 | 
|  |    731 | by (full_simp_tac (simpset() addsimps [real_of_posnat_two,
 | 
|  |    732 |     real_add_assoc]) 1);
 | 
|  |    733 | qed "real_of_nat_one";
 | 
|  |    734 | 
 | 
|  |    735 | Goalw [real_of_nat_def]
 | 
|  |    736 |           "real_of_nat n1 + real_of_nat n2 = real_of_nat (n1 + n2)";
 | 
|  |    737 | by (simp_tac (simpset() addsimps 
 | 
|  |    738 |     [real_of_posnat_add,real_add_assoc RS sym]) 1);
 | 
|  |    739 | qed "real_of_nat_add";
 | 
|  |    740 | 
 | 
|  |    741 | Goalw [real_of_nat_def] "real_of_nat (Suc n) = real_of_nat n + 1r";
 | 
|  |    742 | by (simp_tac (simpset() addsimps [real_of_posnat_Suc] @ real_add_ac) 1);
 | 
|  |    743 | qed "real_of_nat_Suc";
 | 
|  |    744 |     
 | 
|  |    745 | Goalw [real_of_nat_def] "(n < m) = (real_of_nat n < real_of_nat m)";
 | 
|  |    746 | by Auto_tac;
 | 
|  |    747 | qed "real_of_nat_less_iff";
 | 
|  |    748 | 
 | 
|  |    749 | Addsimps [real_of_nat_less_iff RS sym];
 | 
|  |    750 | 
 | 
|  |    751 | Goal "inj real_of_nat";
 | 
|  |    752 | by (rtac injI 1);
 | 
|  |    753 | by (auto_tac (claset() addSIs [inj_real_of_posnat RS injD],
 | 
|  |    754 | 	      simpset() addsimps [real_of_nat_def,real_add_right_cancel]));
 | 
|  |    755 | qed "inj_real_of_nat";
 | 
|  |    756 | 
 | 
|  |    757 | Goalw [real_of_nat_def] "0r <= real_of_nat n";
 | 
|  |    758 | by (res_inst_tac [("C","1r")] real_le_add_right_cancel 1);
 | 
|  |    759 | by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
 | 
|  |    760 | qed "real_of_nat_ge_zero";
 | 
|  |    761 | Addsimps [real_of_nat_ge_zero];
 | 
|  |    762 | 
 | 
|  |    763 | Goal "real_of_nat n1 * real_of_nat n2 = real_of_nat (n1 * n2)";
 | 
|  |    764 | by (induct_tac "n1" 1);
 | 
|  |    765 | by (dtac sym 2);
 | 
|  |    766 | by (auto_tac (claset(),
 | 
|  |    767 | 	      simpset() addsimps [real_of_nat_zero,
 | 
|  |    768 | 				  real_of_nat_add RS sym,real_of_nat_Suc,
 | 
|  |    769 | 				  real_add_mult_distrib, real_add_commute]));
 | 
|  |    770 | qed "real_of_nat_mult";
 | 
|  |    771 | 
 | 
|  |    772 | Goal "(real_of_nat n = real_of_nat m) = (n = m)";
 | 
|  |    773 | by (auto_tac (claset() addDs [inj_real_of_nat RS injD],
 | 
|  |    774 |               simpset()));
 | 
|  |    775 | qed "real_of_nat_eq_cancel";
 | 
|  |    776 | 
 | 
|  |    777 | (*------- lemmas -------*)
 | 
|  |    778 | goal NatDef.thy "!!m. [| m < Suc n; n <= m |] ==> m = n";
 | 
|  |    779 | by (auto_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_asym],
 | 
|  |    780 | 	      simpset() addsimps [less_Suc_eq]));
 | 
|  |    781 | qed "lemma_nat_not_dense";
 | 
|  |    782 | 
 | 
|  |    783 | goal NatDef.thy "!!m. [| m <= Suc n; n < m |] ==> m = Suc n";
 | 
|  |    784 | by (auto_tac (claset() addSDs [le_imp_less_or_eq] addIs [less_asym],
 | 
|  |    785 | 	      simpset() addsimps [le_Suc_eq]));
 | 
|  |    786 | qed "lemma_nat_not_dense2";
 | 
|  |    787 | 
 | 
|  |    788 | goal NatDef.thy "!!m. m < Suc n ==> ~ Suc n <= m";
 | 
|  |    789 | by (blast_tac (claset() addDs [less_le_trans] addIs [less_asym]) 1);
 | 
|  |    790 | qed "lemma_not_leI";
 | 
|  |    791 | 
 | 
|  |    792 | goalw NatDef.thy [le_def] "!!m. ~ Suc n <= m ==> ~ Suc (Suc n) <= m";
 | 
|  |    793 | by Auto_tac;
 | 
|  |    794 | qed "lemma_not_leI2";
 | 
|  |    795 | 
 | 
|  |    796 | (*------- lemmas -------*)
 | 
|  |    797 | val [prem] = goal Arith.thy "n < Suc(m) ==> Suc(m)-n = Suc(m-n)";
 | 
|  |    798 | by (rtac (prem RS rev_mp) 1);
 | 
|  |    799 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
|  |    800 | by (ALLGOALS Asm_simp_tac);
 | 
|  |    801 | qed "Suc_diff_n";
 | 
|  |    802 | 
 | 
|  |    803 | (* Goalw  [real_of_nat_def] 
 | 
|  |    804 |    "real_of_nat (n1 - n2) = real_of_nat n1 + -real_of_nat n2";*)
 | 
|  |    805 | 
 | 
|  |    806 | Goal "n2 < n1 --> real_of_nat (n1 - n2) = real_of_nat n1 + (-real_of_nat n2)";
 | 
|  |    807 | by (induct_tac "n1" 1);
 | 
|  |    808 | by (Step_tac 1 THEN dtac leI 1 THEN dtac sym 2);
 | 
|  |    809 | by (dtac lemma_nat_not_dense 1);
 | 
|  |    810 | by (auto_tac (claset(),
 | 
|  |    811 | 	      simpset() addsimps [real_of_nat_Suc, real_of_nat_zero] @ 
 | 
|  |    812 | 	                         real_add_ac));
 | 
|  |    813 | by (asm_full_simp_tac (simpset() addsimps [real_of_nat_one RS sym,
 | 
|  |    814 | 					   real_of_nat_add,Suc_diff_n]) 1);
 | 
|  |    815 | qed "real_of_nat_minus";
 | 
|  |    816 | 
 | 
|  |    817 | 
 |