| author | oheimb | 
| Thu, 24 Sep 1998 15:36:16 +0200 | |
| changeset 5548 | 5cd3396802f5 | 
| parent 5535 | 678999604ee9 | 
| child 7077 | 60b098bb8b8a | 
| permissions | -rw-r--r-- | 
| 5078 | 1 | (* Title : PRat.ML | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Description : The positive rationals | |
| 5 | *) | |
| 6 | ||
| 7 | open PRat; | |
| 8 | ||
| 9 | Delrules [equalityI]; | |
| 10 | ||
| 11 | (*** Many theorems similar to those in Integ.thy ***) | |
| 12 | (*** Proving that ratrel is an equivalence relation ***) | |
| 13 | ||
| 14 | Goal | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 15 | "[| (x1::pnat) * y2 = x2 * y1; x2 * y3 = x3 * y2 |] \ | 
| 5078 | 16 | \ ==> x1 * y3 = x3 * y1"; | 
| 17 | by (res_inst_tac [("k1","y2")] (pnat_mult_cancel1 RS iffD1) 1);
 | |
| 18 | by (auto_tac (claset(), simpset() addsimps [pnat_mult_assoc RS sym])); | |
| 19 | by (auto_tac (claset(),simpset() addsimps [pnat_mult_commute])); | |
| 20 | by (dres_inst_tac [("s","x2 * y3")] sym 1);
 | |
| 21 | by (asm_simp_tac (simpset() addsimps [pnat_mult_left_commute, | |
| 22 | pnat_mult_commute]) 1); | |
| 23 | qed "prat_trans_lemma"; | |
| 24 | ||
| 25 | (** Natural deduction for ratrel **) | |
| 26 | ||
| 27 | Goalw [ratrel_def] | |
| 28 | "(((x1,y1),(x2,y2)): ratrel) = (x1 * y2 = x2 * y1)"; | |
| 29 | by (Fast_tac 1); | |
| 30 | qed "ratrel_iff"; | |
| 31 | ||
| 32 | Goalw [ratrel_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 33 | "[| x1 * y2 = x2 * y1 |] ==> ((x1,y1),(x2,y2)): ratrel"; | 
| 5078 | 34 | by (Fast_tac 1); | 
| 35 | qed "ratrelI"; | |
| 36 | ||
| 37 | Goalw [ratrel_def] | |
| 38 | "p: ratrel --> (EX x1 y1 x2 y2. \ | |
| 39 | \ p = ((x1,y1),(x2,y2)) & x1 *y2 = x2 *y1)"; | |
| 40 | by (Fast_tac 1); | |
| 41 | qed "ratrelE_lemma"; | |
| 42 | ||
| 43 | val [major,minor] = goal thy | |
| 44 | "[| p: ratrel; \ | |
| 45 | \ !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2)); x1*y2 = x2*y1 \ | |
| 46 | \ |] ==> Q |] ==> Q"; | |
| 47 | by (cut_facts_tac [major RS (ratrelE_lemma RS mp)] 1); | |
| 48 | by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); | |
| 49 | qed "ratrelE"; | |
| 50 | ||
| 51 | AddSIs [ratrelI]; | |
| 52 | AddSEs [ratrelE]; | |
| 53 | ||
| 54 | Goal "(x,x): ratrel"; | |
| 55 | by (stac surjective_pairing 1 THEN rtac (refl RS ratrelI) 1); | |
| 56 | qed "ratrel_refl"; | |
| 57 | ||
| 58 | Goalw [equiv_def, refl_def, sym_def, trans_def] | |
| 59 |     "equiv {x::(pnat*pnat).True} ratrel";
 | |
| 60 | by (fast_tac (claset() addSIs [ratrel_refl] | |
| 61 | addSEs [sym, prat_trans_lemma]) 1); | |
| 62 | qed "equiv_ratrel"; | |
| 63 | ||
| 64 | val equiv_ratrel_iff = | |
| 65 | [TrueI, TrueI] MRS | |
| 66 | ([CollectI, CollectI] MRS | |
| 67 | (equiv_ratrel RS eq_equiv_class_iff)); | |
| 68 | ||
| 69 | Goalw  [prat_def,ratrel_def,quotient_def] "ratrel^^{(x,y)}:prat";
 | |
| 70 | by (Blast_tac 1); | |
| 71 | qed "ratrel_in_prat"; | |
| 72 | ||
| 73 | Goal "inj_on Abs_prat prat"; | |
| 74 | by (rtac inj_on_inverseI 1); | |
| 75 | by (etac Abs_prat_inverse 1); | |
| 76 | qed "inj_on_Abs_prat"; | |
| 77 | ||
| 78 | Addsimps [equiv_ratrel_iff,inj_on_Abs_prat RS inj_on_iff, | |
| 79 | ratrel_iff, ratrel_in_prat, Abs_prat_inverse]; | |
| 80 | ||
| 81 | Addsimps [equiv_ratrel RS eq_equiv_class_iff]; | |
| 82 | val eq_ratrelD = equiv_ratrel RSN (2,eq_equiv_class); | |
| 83 | ||
| 84 | Goal "inj(Rep_prat)"; | |
| 85 | by (rtac inj_inverseI 1); | |
| 86 | by (rtac Rep_prat_inverse 1); | |
| 87 | qed "inj_Rep_prat"; | |
| 88 | ||
| 89 | (** prat_pnat: the injection from pnat to prat **) | |
| 90 | Goal "inj(prat_pnat)"; | |
| 91 | by (rtac injI 1); | |
| 92 | by (rewtac prat_pnat_def); | |
| 93 | by (dtac (inj_on_Abs_prat RS inj_onD) 1); | |
| 94 | by (REPEAT (rtac ratrel_in_prat 1)); | |
| 95 | by (dtac eq_equiv_class 1); | |
| 96 | by (rtac equiv_ratrel 1); | |
| 97 | by (Fast_tac 1); | |
| 98 | by Safe_tac; | |
| 99 | by (Asm_full_simp_tac 1); | |
| 100 | qed "inj_prat_pnat"; | |
| 101 | ||
| 102 | (* lcp's original eq_Abs_Integ *) | |
| 103 | val [prem] = goal thy | |
| 104 |     "(!!x y. z = Abs_prat(ratrel^^{(x,y)}) ==> P) ==> P";
 | |
| 105 | by (res_inst_tac [("x1","z")] 
 | |
| 106 | (rewrite_rule [prat_def] Rep_prat RS quotientE) 1); | |
| 107 | by (dres_inst_tac [("f","Abs_prat")] arg_cong 1);
 | |
| 108 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 109 | by (rtac prem 1); | |
| 110 | by (asm_full_simp_tac (simpset() addsimps [Rep_prat_inverse]) 1); | |
| 111 | qed "eq_Abs_prat"; | |
| 112 | ||
| 113 | (**** qinv: inverse on prat ****) | |
| 114 | ||
| 115 | Goalw [congruent_def] | |
| 116 |   "congruent ratrel (%p. split (%x y. ratrel^^{(y,x)}) p)";
 | |
| 117 | by Safe_tac; | |
| 118 | by (asm_full_simp_tac (simpset() addsimps [pnat_mult_commute]) 1); | |
| 119 | qed "qinv_congruent"; | |
| 120 | ||
| 121 | (*Resolve th against the corresponding facts for qinv*) | |
| 122 | val qinv_ize = RSLIST [equiv_ratrel, qinv_congruent]; | |
| 123 | ||
| 124 | Goalw [qinv_def] | |
| 125 |       "qinv (Abs_prat(ratrel^^{(x,y)})) = Abs_prat(ratrel ^^ {(y,x)})";
 | |
| 126 | by (res_inst_tac [("f","Abs_prat")] arg_cong 1);
 | |
| 127 | by (simp_tac (simpset() addsimps | |
| 128 | [ratrel_in_prat RS Abs_prat_inverse,qinv_ize UN_equiv_class]) 1); | |
| 129 | qed "qinv"; | |
| 130 | ||
| 131 | Goal "qinv (qinv z) = z"; | |
| 132 | by (res_inst_tac [("z","z")] eq_Abs_prat 1);
 | |
| 133 | by (asm_simp_tac (simpset() addsimps [qinv]) 1); | |
| 134 | qed "qinv_qinv"; | |
| 135 | ||
| 136 | Goal "inj(qinv)"; | |
| 137 | by (rtac injI 1); | |
| 138 | by (dres_inst_tac [("f","qinv")] arg_cong 1);
 | |
| 139 | by (asm_full_simp_tac (simpset() addsimps [qinv_qinv]) 1); | |
| 140 | qed "inj_qinv"; | |
| 141 | ||
| 142 | Goalw [prat_pnat_def] "qinv($# (Abs_pnat 1)) = $#(Abs_pnat 1)"; | |
| 143 | by (simp_tac (simpset() addsimps [qinv]) 1); | |
| 144 | qed "qinv_1"; | |
| 145 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 146 | Goal "!!(x1::pnat). [| x1 * y2 = x2 * y1 |] ==> \ | 
| 5078 | 147 | \ (x * y1 + x1 * ya) * (ya * y2) = (x * y2 + x2 * ya) * (ya * y1)"; | 
| 148 | by (auto_tac (claset() addSIs [pnat_same_multI2], | |
| 149 | simpset() addsimps [pnat_add_mult_distrib, | |
| 150 | pnat_mult_assoc])); | |
| 151 | by (res_inst_tac [("n1","y2")] (pnat_mult_commute RS subst) 1);
 | |
| 152 | by (auto_tac (claset() addIs [pnat_add_left_cancel RS iffD2],simpset() addsimps pnat_mult_ac)); | |
| 153 | by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS subst) 1);
 | |
| 154 | by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS ssubst) 1);
 | |
| 155 | by (auto_tac (claset(),simpset() addsimps [pnat_mult_assoc RS sym])); | |
| 156 | qed "prat_add_congruent2_lemma"; | |
| 157 | ||
| 158 | Goal | |
| 159 | "congruent2 ratrel (%p1 p2. \ | |
| 160 | \         split (%x1 y1. split (%x2 y2. ratrel^^{(x1*y2 + x2*y1, y1*y2)}) p2) p1)";
 | |
| 161 | by (rtac (equiv_ratrel RS congruent2_commuteI) 1); | |
| 162 | by Safe_tac; | |
| 163 | by (rewtac split_def); | |
| 164 | by (asm_simp_tac (simpset() addsimps [pnat_mult_commute,pnat_add_commute]) 1); | |
| 165 | by (auto_tac (claset(),simpset() addsimps [prat_add_congruent2_lemma])); | |
| 166 | qed "prat_add_congruent2"; | |
| 167 | ||
| 168 | (*Resolve th against the corresponding facts for prat_add*) | |
| 169 | val prat_add_ize = RSLIST [equiv_ratrel, prat_add_congruent2]; | |
| 170 | ||
| 171 | Goalw [prat_add_def] | |
| 172 |    "Abs_prat((ratrel^^{(x1,y1)})) + Abs_prat((ratrel^^{(x2,y2)})) =   \
 | |
| 173 | \   Abs_prat(ratrel ^^ {(x1*y2 + x2*y1, y1*y2)})";
 | |
| 174 | by (simp_tac (simpset() addsimps [prat_add_ize UN_equiv_class2]) 1); | |
| 175 | qed "prat_add"; | |
| 176 | ||
| 177 | Goal "(z::prat) + w = w + z"; | |
| 178 | by (res_inst_tac [("z","z")] eq_Abs_prat 1);
 | |
| 179 | by (res_inst_tac [("z","w")] eq_Abs_prat 1);
 | |
| 5535 | 180 | by (asm_simp_tac | 
| 181 | (simpset() addsimps [prat_add] @ pnat_add_ac @ pnat_mult_ac) 1); | |
| 5078 | 182 | qed "prat_add_commute"; | 
| 183 | ||
| 184 | Goal "((z1::prat) + z2) + z3 = z1 + (z2 + z3)"; | |
| 185 | by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
 | |
| 186 | by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
 | |
| 187 | by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
 | |
| 5535 | 188 | by (asm_simp_tac (simpset() addsimps [pnat_add_mult_distrib2,prat_add] @ | 
| 189 | pnat_add_ac @ pnat_mult_ac) 1); | |
| 5078 | 190 | qed "prat_add_assoc"; | 
| 191 | ||
| 192 | qed_goal "prat_add_left_commute" thy | |
| 193 | "(z1::prat) + (z2 + z3) = z2 + (z1 + z3)" | |
| 194 | (fn _ => [rtac (prat_add_commute RS trans) 1, rtac (prat_add_assoc RS trans) 1, | |
| 195 | rtac (prat_add_commute RS arg_cong) 1]); | |
| 196 | ||
| 197 | (* Positive Rational addition is an AC operator *) | |
| 198 | val prat_add_ac = [prat_add_assoc, prat_add_commute, prat_add_left_commute]; | |
| 199 | ||
| 200 | ||
| 201 | (*** Congruence property for multiplication ***) | |
| 202 | ||
| 203 | Goalw [congruent2_def] | |
| 204 | "congruent2 ratrel (%p1 p2. \ | |
| 205 | \         split (%x1 y1. split (%x2 y2. ratrel^^{(x1*x2, y1*y2)}) p2) p1)";
 | |
| 206 | ||
| 207 | (*Proof via congruent2_commuteI seems longer*) | |
| 208 | by Safe_tac; | |
| 209 | by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc]) 1); | |
| 210 | (*The rest should be trivial, but rearranging terms is hard*) | |
| 211 | by (res_inst_tac [("x1","x1a")] (pnat_mult_left_commute RS ssubst) 1);
 | |
| 212 | by (asm_simp_tac (simpset() addsimps [pnat_mult_assoc RS sym]) 1); | |
| 213 | by (asm_simp_tac (simpset() addsimps pnat_mult_ac) 1); | |
| 214 | qed "pnat_mult_congruent2"; | |
| 215 | ||
| 216 | (*Resolve th against the corresponding facts for pnat_mult*) | |
| 217 | val prat_mult_ize = RSLIST [equiv_ratrel, pnat_mult_congruent2]; | |
| 218 | ||
| 219 | Goalw [prat_mult_def] | |
| 220 |   "Abs_prat(ratrel^^{(x1,y1)}) * Abs_prat(ratrel^^{(x2,y2)}) = \
 | |
| 221 | \  Abs_prat(ratrel^^{(x1*x2, y1*y2)})";
 | |
| 222 | by (asm_simp_tac | |
| 223 | (simpset() addsimps [prat_mult_ize UN_equiv_class2]) 1); | |
| 224 | qed "prat_mult"; | |
| 225 | ||
| 226 | Goal "(z::prat) * w = w * z"; | |
| 227 | by (res_inst_tac [("z","z")] eq_Abs_prat 1);
 | |
| 228 | by (res_inst_tac [("z","w")] eq_Abs_prat 1);
 | |
| 5535 | 229 | by (asm_simp_tac (simpset() addsimps pnat_mult_ac @ [prat_mult]) 1); | 
| 5078 | 230 | qed "prat_mult_commute"; | 
| 231 | ||
| 232 | Goal "((z1::prat) * z2) * z3 = z1 * (z2 * z3)"; | |
| 233 | by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
 | |
| 234 | by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
 | |
| 235 | by (res_inst_tac [("z","z3")] eq_Abs_prat 1);
 | |
| 236 | by (asm_simp_tac (simpset() addsimps [prat_mult, pnat_mult_assoc]) 1); | |
| 237 | qed "prat_mult_assoc"; | |
| 238 | ||
| 239 | (*For AC rewriting*) | |
| 240 | Goal "(x::prat)*(y*z)=y*(x*z)"; | |
| 241 | by (rtac (prat_mult_commute RS trans) 1); | |
| 242 | by (rtac (prat_mult_assoc RS trans) 1); | |
| 243 | by (rtac (prat_mult_commute RS arg_cong) 1); | |
| 244 | qed "prat_mult_left_commute"; | |
| 245 | ||
| 246 | (*Positive Rational multiplication is an AC operator*) | |
| 247 | val prat_mult_ac = [prat_mult_assoc,prat_mult_commute,prat_mult_left_commute]; | |
| 248 | ||
| 249 | Goalw [prat_pnat_def] "($#Abs_pnat 1) * z = z"; | |
| 250 | by (res_inst_tac [("z","z")] eq_Abs_prat 1);
 | |
| 251 | by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1); | |
| 252 | qed "prat_mult_1"; | |
| 253 | ||
| 254 | Goalw [prat_pnat_def] "z * ($#Abs_pnat 1) = z"; | |
| 255 | by (res_inst_tac [("z","z")] eq_Abs_prat 1);
 | |
| 256 | by (asm_full_simp_tac (simpset() addsimps [prat_mult] @ pnat_mult_ac) 1); | |
| 257 | qed "prat_mult_1_right"; | |
| 258 | ||
| 259 | Goalw [prat_pnat_def] | |
| 260 | "$#((z1::pnat) + z2) = $#z1 + $#z2"; | |
| 261 | by (asm_simp_tac (simpset() addsimps [prat_add, | |
| 262 | pnat_add_mult_distrib,pnat_mult_1]) 1); | |
| 263 | qed "prat_pnat_add"; | |
| 264 | ||
| 265 | Goalw [prat_pnat_def] | |
| 266 | "$#((z1::pnat) * z2) = $#z1 * $#z2"; | |
| 267 | by (asm_simp_tac (simpset() addsimps [prat_mult, | |
| 268 | pnat_mult_1]) 1); | |
| 269 | qed "prat_pnat_mult"; | |
| 270 | ||
| 271 | (*** prat_mult and qinv ***) | |
| 272 | ||
| 273 | Goalw [prat_def,prat_pnat_def] "qinv (q) * q = $# (Abs_pnat 1)"; | |
| 274 | by (res_inst_tac [("z","q")] eq_Abs_prat 1);
 | |
| 275 | by (asm_full_simp_tac (simpset() addsimps [qinv, | |
| 276 | prat_mult,pnat_mult_1,pnat_mult_1_left, | |
| 277 | pnat_mult_commute]) 1); | |
| 278 | qed "prat_mult_qinv"; | |
| 279 | ||
| 280 | Goal "q * qinv (q) = $# (Abs_pnat 1)"; | |
| 281 | by (rtac (prat_mult_commute RS subst) 1); | |
| 282 | by (simp_tac (simpset() addsimps [prat_mult_qinv]) 1); | |
| 283 | qed "prat_mult_qinv_right"; | |
| 284 | ||
| 285 | Goal "? y. (x::prat) * y = $# Abs_pnat 1"; | |
| 286 | by (fast_tac (claset() addIs [prat_mult_qinv_right]) 1); | |
| 287 | qed "prat_qinv_ex"; | |
| 288 | ||
| 289 | Goal "?! y. (x::prat) * y = $# Abs_pnat 1"; | |
| 290 | by (auto_tac (claset() addIs [prat_mult_qinv_right],simpset())); | |
| 291 | by (dres_inst_tac [("f","%x. ya*x")] arg_cong 1);
 | |
| 292 | by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc RS sym]) 1); | |
| 293 | by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute, | |
| 294 | prat_mult_1,prat_mult_1_right]) 1); | |
| 295 | qed "prat_qinv_ex1"; | |
| 296 | ||
| 297 | Goal "?! y. y * (x::prat) = $# Abs_pnat 1"; | |
| 298 | by (auto_tac (claset() addIs [prat_mult_qinv],simpset())); | |
| 299 | by (dres_inst_tac [("f","%x. x*ya")] arg_cong 1);
 | |
| 300 | by (asm_full_simp_tac (simpset() addsimps [prat_mult_assoc]) 1); | |
| 301 | by (asm_full_simp_tac (simpset() addsimps [prat_mult_commute, | |
| 302 | prat_mult_1,prat_mult_1_right]) 1); | |
| 303 | qed "prat_qinv_left_ex1"; | |
| 304 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 305 | Goal "x * y = $# Abs_pnat 1 ==> x = qinv y"; | 
| 5078 | 306 | by (cut_inst_tac [("q","y")] prat_mult_qinv 1);
 | 
| 307 | by (res_inst_tac [("x1","y")] (prat_qinv_left_ex1 RS ex1E) 1);
 | |
| 308 | by (Blast_tac 1); | |
| 309 | qed "prat_mult_inv_qinv"; | |
| 310 | ||
| 311 | Goal "? y. x = qinv y"; | |
| 312 | by (cut_inst_tac [("x","x")] prat_qinv_ex 1);
 | |
| 313 | by (etac exE 1 THEN dtac prat_mult_inv_qinv 1); | |
| 314 | by (Fast_tac 1); | |
| 315 | qed "prat_as_inverse_ex"; | |
| 316 | ||
| 317 | Goal "qinv(x*y) = qinv(x)*qinv(y)"; | |
| 318 | by (res_inst_tac [("z","x")] eq_Abs_prat 1);
 | |
| 319 | by (res_inst_tac [("z","y")] eq_Abs_prat 1);
 | |
| 320 | by (auto_tac (claset(),simpset() addsimps [qinv,prat_mult])); | |
| 321 | qed "qinv_mult_eq"; | |
| 322 | ||
| 323 | (** Lemmas **) | |
| 324 | ||
| 325 | qed_goal "prat_add_assoc_cong" thy | |
| 326 | "!!z. (z::prat) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" | |
| 327 | (fn _ => [(asm_simp_tac (simpset() addsimps [prat_add_assoc RS sym]) 1)]); | |
| 328 | ||
| 329 | qed_goal "prat_add_assoc_swap" thy "(z::prat) + (v + w) = v + (z + w)" | |
| 330 | (fn _ => [(REPEAT (ares_tac [prat_add_commute RS prat_add_assoc_cong] 1))]); | |
| 331 | ||
| 332 | Goal "((z1::prat) + z2) * w = (z1 * w) + (z2 * w)"; | |
| 333 | by (res_inst_tac [("z","z1")] eq_Abs_prat 1);
 | |
| 334 | by (res_inst_tac [("z","z2")] eq_Abs_prat 1);
 | |
| 335 | by (res_inst_tac [("z","w")] eq_Abs_prat 1);
 | |
| 336 | by (asm_simp_tac | |
| 5535 | 337 | (simpset() addsimps [pnat_add_mult_distrib2, prat_add, prat_mult] @ | 
| 338 | pnat_add_ac @ pnat_mult_ac) 1); | |
| 5078 | 339 | qed "prat_add_mult_distrib"; | 
| 340 | ||
| 341 | val prat_mult_commute'= read_instantiate [("z","w")] prat_mult_commute;
 | |
| 342 | ||
| 343 | Goal "(w::prat) * (z1 + z2) = (w * z1) + (w * z2)"; | |
| 344 | by (simp_tac (simpset() addsimps [prat_mult_commute',prat_add_mult_distrib]) 1); | |
| 345 | qed "prat_add_mult_distrib2"; | |
| 346 | ||
| 347 | val prat_mult_simps = [prat_mult_1, prat_mult_1_right, | |
| 348 | prat_mult_qinv, prat_mult_qinv_right]; | |
| 349 | Addsimps prat_mult_simps; | |
| 350 | ||
| 351 | (*** theorems for ordering ***) | |
| 352 | (* prove introduction and elimination rules for prat_less *) | |
| 353 | ||
| 354 | Goalw [prat_less_def] | |
| 355 | "Q1 < (Q2::prat) = (EX Q3. Q1 + Q3 = Q2)"; | |
| 356 | by (Fast_tac 1); | |
| 357 | qed "prat_less_iff"; | |
| 358 | ||
| 359 | Goalw [prat_less_def] | |
| 360 | "!!(Q1::prat). Q1 + Q3 = Q2 ==> Q1 < Q2"; | |
| 361 | by (Fast_tac 1); | |
| 362 | qed "prat_lessI"; | |
| 363 | ||
| 364 | (* ordering on positive fractions in terms of existence of sum *) | |
| 365 | Goalw [prat_less_def] | |
| 366 | "Q1 < (Q2::prat) --> (EX Q3. Q1 + Q3 = Q2)"; | |
| 367 | by (Fast_tac 1); | |
| 368 | qed "prat_lessE_lemma"; | |
| 369 | ||
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 370 | Goal "!!P. [| Q1 < (Q2::prat); \ | 
| 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 371 | \ !! (Q3::prat). Q1 + Q3 = Q2 ==> P |] \ | 
| 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 372 | \ ==> P"; | 
| 5078 | 373 | by (dtac (prat_lessE_lemma RS mp) 1); | 
| 374 | by Auto_tac; | |
| 375 | qed "prat_lessE"; | |
| 376 | ||
| 377 | (* qless is a strong order i.e nonreflexive and transitive *) | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 378 | Goal "!!(q1::prat). [| q1 < q2; q2 < q3 |] ==> q1 < q3"; | 
| 5078 | 379 | by (REPEAT(dtac (prat_lessE_lemma RS mp) 1)); | 
| 380 | by (REPEAT(etac exE 1)); | |
| 381 | by (hyp_subst_tac 1); | |
| 382 | by (res_inst_tac [("Q3.0","Q3 + Q3a")] prat_lessI 1);
 | |
| 383 | by (auto_tac (claset(),simpset() addsimps [prat_add_assoc])); | |
| 384 | qed "prat_less_trans"; | |
| 385 | ||
| 386 | Goal "~q < (q::prat)"; | |
| 387 | by (EVERY1[rtac notI, dtac (prat_lessE_lemma RS mp)]); | |
| 388 | by (res_inst_tac [("z","q")] eq_Abs_prat 1);
 | |
| 389 | by (res_inst_tac [("z","Q3")] eq_Abs_prat 1);
 | |
| 390 | by (etac exE 1 THEN res_inst_tac [("z","Q3a")] eq_Abs_prat 1);
 | |
| 391 | by (REPEAT(hyp_subst_tac 1)); | |
| 392 | by (asm_full_simp_tac (simpset() addsimps [prat_add, | |
| 393 | pnat_no_add_ident,pnat_add_mult_distrib2] @ pnat_mult_ac) 1); | |
| 394 | qed "prat_less_not_refl"; | |
| 395 | ||
| 396 | (*** y < y ==> P ***) | |
| 397 | bind_thm("prat_less_irrefl",prat_less_not_refl RS notE);
 | |
| 398 | ||
| 5459 | 399 | Goal "!! (q1::prat). q1 < q2 ==> ~ q2 < q1"; | 
| 400 | by (rtac notI 1); | |
| 5078 | 401 | by (dtac prat_less_trans 1 THEN assume_tac 1); | 
| 402 | by (asm_full_simp_tac (simpset() addsimps [prat_less_not_refl]) 1); | |
| 5459 | 403 | qed "prat_less_not_sym"; | 
| 5078 | 404 | |
| 5459 | 405 | (* [| x < y; ~P ==> y < x |] ==> P *) | 
| 406 | bind_thm ("prat_less_asym", prat_less_not_sym RS swap);
 | |
| 5078 | 407 | |
| 408 | (* half of positive fraction exists- Gleason p. 120- Proposition 9-2.6(i)*) | |
| 409 | Goal "!(q::prat). ? x. x + x = q"; | |
| 410 | by (rtac allI 1); | |
| 411 | by (res_inst_tac [("z","q")] eq_Abs_prat 1);
 | |
| 412 | by (res_inst_tac [("x","Abs_prat (ratrel ^^ {(x, y+y)})")] exI 1);
 | |
| 413 | by (auto_tac (claset(),simpset() addsimps | |
| 414 | [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib, | |
| 415 | pnat_add_mult_distrib2])); | |
| 416 | qed "lemma_prat_dense"; | |
| 417 | ||
| 418 | Goal "? (x::prat). x + x = q"; | |
| 419 | by (res_inst_tac [("z","q")] eq_Abs_prat 1);
 | |
| 420 | by (res_inst_tac [("x","Abs_prat (ratrel ^^ {(x, y+y)})")] exI 1);
 | |
| 421 | by (auto_tac (claset(),simpset() addsimps | |
| 422 | [prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib, | |
| 423 | pnat_add_mult_distrib2])); | |
| 424 | qed "prat_lemma_dense"; | |
| 425 | ||
| 426 | (* there exists a number between any two positive fractions *) | |
| 427 | (* Gleason p. 120- Proposition 9-2.6(iv) *) | |
| 428 | Goalw [prat_less_def] | |
| 429 | "!! (q1::prat). q1 < q2 ==> ? x. q1 < x & x < q2"; | |
| 430 | by (auto_tac (claset() addIs [lemma_prat_dense],simpset())); | |
| 431 | by (res_inst_tac [("x","T")] (lemma_prat_dense RS allE) 1);
 | |
| 432 | by (etac exE 1); | |
| 433 | by (res_inst_tac [("x","q1 + x")] exI 1);
 | |
| 434 | by (auto_tac (claset() addIs [prat_lemma_dense], | |
| 435 | simpset() addsimps [prat_add_assoc])); | |
| 436 | qed "prat_dense"; | |
| 437 | ||
| 438 | (* ordering of addition for positive fractions *) | |
| 439 | Goalw [prat_less_def] | |
| 440 | "!!(q1::prat). q1 < q2 ==> q1 + x < q2 + x"; | |
| 441 | by (Step_tac 1); | |
| 442 | by (res_inst_tac [("x","T")] exI 1);
 | |
| 443 | by (auto_tac (claset(),simpset() addsimps prat_add_ac)); | |
| 444 | qed "prat_add_less2_mono1"; | |
| 445 | ||
| 446 | Goal | |
| 447 | "!!(q1::prat). q1 < q2 ==> x + q1 < x + q2"; | |
| 448 | by (auto_tac (claset() addIs [prat_add_less2_mono1], | |
| 449 | simpset() addsimps [prat_add_commute])); | |
| 450 | qed "prat_add_less2_mono2"; | |
| 451 | ||
| 452 | (* ordering of multiplication for positive fractions *) | |
| 453 | Goalw [prat_less_def] | |
| 454 | "!!(q1::prat). q1 < q2 ==> q1 * x < q2 * x"; | |
| 455 | by (Step_tac 1); | |
| 456 | by (res_inst_tac [("x","T*x")] exI 1);
 | |
| 457 | by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib])); | |
| 458 | qed "prat_mult_less2_mono1"; | |
| 459 | ||
| 460 | Goal "!!(q1::prat). q1 < q2 ==> x * q1 < x * q2"; | |
| 461 | by (auto_tac (claset() addDs [prat_mult_less2_mono1], | |
| 462 | simpset() addsimps [prat_mult_commute])); | |
| 463 | qed "prat_mult_left_less2_mono1"; | |
| 464 | ||
| 465 | (* there is no smallest positive fraction *) | |
| 466 | Goalw [prat_less_def] "? (x::prat). x < y"; | |
| 467 | by (cut_facts_tac [lemma_prat_dense] 1); | |
| 468 | by (Fast_tac 1); | |
| 469 | qed "qless_Ex"; | |
| 470 | ||
| 471 | (* lemma for proving $< is linear *) | |
| 472 | Goalw [prat_def,prat_less_def] | |
| 473 |       "ratrel ^^ {(x, y * ya)} : {p::(pnat*pnat).True}/ratrel";
 | |
| 474 | by (asm_full_simp_tac (simpset() addsimps [ratrel_def,quotient_def]) 1); | |
| 475 | by (Blast_tac 1); | |
| 476 | qed "lemma_prat_less_linear"; | |
| 477 | ||
| 478 | (* linearity of < -- Gleason p. 120 - Proposition 9-2.6 *) | |
| 479 | (*** FIXME Proof long ***) | |
| 480 | Goalw [prat_less_def] | |
| 481 | "(q1::prat) < q2 | q1 = q2 | q2 < q1"; | |
| 482 | by (res_inst_tac [("z","q1")] eq_Abs_prat 1);
 | |
| 483 | by (res_inst_tac [("z","q2")] eq_Abs_prat 1);
 | |
| 484 | by (Step_tac 1 THEN REPEAT(dtac (not_ex RS iffD1) 1) | |
| 485 | THEN Auto_tac); | |
| 486 | by (cut_inst_tac  [("z1.0","x*ya"),
 | |
| 487 |    ("z2.0","xa*y")] pnat_linear_Ex_eq 1);
 | |
| 488 | by (EVERY1[etac disjE,etac exE]); | |
| 489 | by (eres_inst_tac | |
| 490 |     [("x","Abs_prat(ratrel^^{(xb,ya*y)})")] allE 1);
 | |
| 491 | by (asm_full_simp_tac | |
| 492 | (simpset() addsimps [prat_add, pnat_mult_assoc | |
| 493 | RS sym,pnat_add_mult_distrib RS sym]) 1); | |
| 494 | by (EVERY1[asm_full_simp_tac (simpset() addsimps pnat_mult_ac), | |
| 495 | etac disjE, assume_tac, etac exE]); | |
| 496 | by (thin_tac "!T. Abs_prat (ratrel ^^ {(x, y)}) + T ~= \
 | |
| 497 | \     Abs_prat (ratrel ^^ {(xa, ya)})" 1);
 | |
| 498 | by (eres_inst_tac [("x","Abs_prat(ratrel^^{(xb,y*ya)})")] allE 1);
 | |
| 499 | by (asm_full_simp_tac (simpset() addsimps [prat_add, | |
| 500 | pnat_mult_assoc RS sym,pnat_add_mult_distrib RS sym]) 1); | |
| 501 | by (asm_full_simp_tac (simpset() addsimps pnat_mult_ac) 1); | |
| 502 | qed "prat_linear"; | |
| 503 | ||
| 504 | Goal | |
| 505 | "!!(q1::prat). [| q1 < q2 ==> P; q1 = q2 ==> P; \ | |
| 506 | \ q2 < q1 ==> P |] ==> P"; | |
| 507 | by (cut_inst_tac [("q1.0","q1"),("q2.0","q2")] prat_linear 1);
 | |
| 508 | by Auto_tac; | |
| 509 | qed "prat_linear_less2"; | |
| 510 | ||
| 511 | (* Gleason p. 120 -- 9-2.6 (iv) *) | |
| 512 | Goal | |
| 513 | "!!(q1::prat). [| q1 < q2; qinv(q1) = qinv(q2) |] ==> P"; | |
| 514 | by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"),
 | |
| 515 |    ("q2.0","q2")] prat_mult_less2_mono1 1);
 | |
| 516 | by (assume_tac 1); | |
| 517 | by (Asm_full_simp_tac 1 THEN dtac sym 1); | |
| 518 | by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl])); | |
| 519 | qed "lemma1_qinv_prat_less"; | |
| 520 | ||
| 521 | Goal | |
| 522 | "!!(q1::prat). [| q1 < q2; qinv(q1) < qinv(q2) |] ==> P"; | |
| 523 | by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"),
 | |
| 524 |    ("q2.0","q2")] prat_mult_less2_mono1 1);
 | |
| 525 | by (assume_tac 1); | |
| 526 | by (cut_inst_tac [("x","q1"),("q1.0","qinv (q1)"),
 | |
| 527 |    ("q2.0","qinv (q2)")] prat_mult_left_less2_mono1 1);
 | |
| 528 | by Auto_tac; | |
| 529 | by (dres_inst_tac [("q2.0","$#Abs_pnat 1")] prat_less_trans 1);
 | |
| 530 | by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl])); | |
| 531 | qed "lemma2_qinv_prat_less"; | |
| 532 | ||
| 533 | Goal | |
| 534 | "!!(q1::prat). q1 < q2 ==> qinv (q2) < qinv (q1)"; | |
| 535 | by (res_inst_tac [("q2.0","qinv q1"),
 | |
| 536 |          ("q1.0","qinv q2")] prat_linear_less2 1);
 | |
| 537 | by (auto_tac (claset() addEs [lemma1_qinv_prat_less, | |
| 538 | lemma2_qinv_prat_less],simpset())); | |
| 539 | qed "qinv_prat_less"; | |
| 540 | ||
| 541 | Goal "!!(q1::prat). q1 < $#Abs_pnat 1 ==> $#Abs_pnat 1 < qinv(q1)"; | |
| 542 | by (dtac qinv_prat_less 1); | |
| 543 | by (full_simp_tac (simpset() addsimps [qinv_1]) 1); | |
| 544 | qed "prat_qinv_gt_1"; | |
| 545 | ||
| 546 | Goalw [pnat_one_def] "!!(q1::prat). q1 < $#1p ==> $#1p < qinv(q1)"; | |
| 547 | by (etac prat_qinv_gt_1 1); | |
| 548 | qed "prat_qinv_is_gt_1"; | |
| 549 | ||
| 550 | Goalw [prat_less_def] "$#Abs_pnat 1 < $#Abs_pnat 1 + $#Abs_pnat 1"; | |
| 551 | by (Fast_tac 1); | |
| 552 | qed "prat_less_1_2"; | |
| 553 | ||
| 554 | Goal "qinv($#Abs_pnat 1 + $#Abs_pnat 1) < $#Abs_pnat 1"; | |
| 555 | by (cut_facts_tac [prat_less_1_2 RS qinv_prat_less] 1); | |
| 556 | by (asm_full_simp_tac (simpset() addsimps [qinv_1]) 1); | |
| 557 | qed "prat_less_qinv_2_1"; | |
| 558 | ||
| 559 | Goal "!!(x::prat). x < y ==> x*qinv(y) < $#Abs_pnat 1"; | |
| 560 | by (dres_inst_tac [("x","qinv(y)")] prat_mult_less2_mono1 1);
 | |
| 561 | by (Asm_full_simp_tac 1); | |
| 562 | qed "prat_mult_qinv_less_1"; | |
| 563 | ||
| 564 | Goal "(x::prat) < x + x"; | |
| 565 | by (cut_inst_tac [("x","x")] 
 | |
| 566 | (prat_less_1_2 RS prat_mult_left_less2_mono1) 1); | |
| 567 | by (asm_full_simp_tac (simpset() addsimps | |
| 568 | [prat_add_mult_distrib2]) 1); | |
| 569 | qed "prat_self_less_add_self"; | |
| 570 | ||
| 571 | Goalw [prat_less_def] "(x::prat) < y + x"; | |
| 572 | by (res_inst_tac [("x","y")] exI 1);
 | |
| 573 | by (simp_tac (simpset() addsimps [prat_add_commute]) 1); | |
| 574 | qed "prat_self_less_add_right"; | |
| 575 | ||
| 576 | Goal "(x::prat) < x + y"; | |
| 577 | by (rtac (prat_add_commute RS subst) 1); | |
| 578 | by (simp_tac (simpset() addsimps [prat_self_less_add_right]) 1); | |
| 579 | qed "prat_self_less_add_left"; | |
| 580 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 581 | Goalw [prat_less_def] "$#1p < y ==> (x::prat) < x * y"; | 
| 5078 | 582 | by (auto_tac (claset(),simpset() addsimps [pnat_one_def, | 
| 583 | prat_add_mult_distrib2])); | |
| 584 | qed "prat_self_less_mult_right"; | |
| 585 | ||
| 586 | (*** Properties of <= ***) | |
| 587 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 588 | Goalw [prat_le_def] "~(w < z) ==> z <= (w::prat)"; | 
| 5078 | 589 | by (assume_tac 1); | 
| 590 | qed "prat_leI"; | |
| 591 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 592 | Goalw [prat_le_def] "z<=w ==> ~(w<(z::prat))"; | 
| 5078 | 593 | by (assume_tac 1); | 
| 594 | qed "prat_leD"; | |
| 595 | ||
| 596 | val prat_leE = make_elim prat_leD; | |
| 597 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 598 | Goal "(~(w < z)) = (z <= (w::prat))"; | 
| 5078 | 599 | by (fast_tac (claset() addSIs [prat_leI,prat_leD]) 1); | 
| 600 | qed "prat_less_le_iff"; | |
| 601 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 602 | Goalw [prat_le_def] "~ z <= w ==> w<(z::prat)"; | 
| 5078 | 603 | by (Fast_tac 1); | 
| 604 | qed "not_prat_leE"; | |
| 605 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 606 | Goalw [prat_le_def] "z < w ==> z <= (w::prat)"; | 
| 5078 | 607 | by (fast_tac (claset() addEs [prat_less_asym]) 1); | 
| 608 | qed "prat_less_imp_le"; | |
| 609 | ||
| 610 | Goalw [prat_le_def] "!!(x::prat). x <= y ==> x < y | x = y"; | |
| 611 | by (cut_facts_tac [prat_linear] 1); | |
| 612 | by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1); | |
| 613 | qed "prat_le_imp_less_or_eq"; | |
| 614 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 615 | Goalw [prat_le_def] "z<w | z=w ==> z <=(w::prat)"; | 
| 5078 | 616 | by (cut_facts_tac [prat_linear] 1); | 
| 617 | by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1); | |
| 618 | qed "prat_less_or_eq_imp_le"; | |
| 619 | ||
| 620 | Goal "(x <= (y::prat)) = (x < y | x=y)"; | |
| 621 | by (REPEAT(ares_tac [iffI, prat_less_or_eq_imp_le, prat_le_imp_less_or_eq] 1)); | |
| 622 | qed "prat_le_eq_less_or_eq"; | |
| 623 | ||
| 624 | Goal "w <= (w::prat)"; | |
| 625 | by (simp_tac (simpset() addsimps [prat_le_eq_less_or_eq]) 1); | |
| 626 | qed "prat_le_refl"; | |
| 627 | ||
| 628 | val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::prat)"; | |
| 629 | by (dtac prat_le_imp_less_or_eq 1); | |
| 630 | by (fast_tac (claset() addIs [prat_less_trans]) 1); | |
| 631 | qed "prat_le_less_trans"; | |
| 632 | ||
| 633 | Goal "!! (i::prat). [| i < j; j <= k |] ==> i < k"; | |
| 634 | by (dtac prat_le_imp_less_or_eq 1); | |
| 635 | by (fast_tac (claset() addIs [prat_less_trans]) 1); | |
| 636 | qed "prat_less_le_trans"; | |
| 637 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 638 | Goal "[| i <= j; j <= k |] ==> i <= (k::prat)"; | 
| 5078 | 639 | by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq, | 
| 640 | rtac prat_less_or_eq_imp_le, fast_tac (claset() addIs [prat_less_trans])]); | |
| 641 | qed "prat_le_trans"; | |
| 642 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 643 | Goal "[| z <= w; w <= z |] ==> z = (w::prat)"; | 
| 5078 | 644 | by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq, | 
| 645 | fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym])]); | |
| 646 | qed "prat_le_anti_sym"; | |
| 647 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 648 | Goal "[| ~ y < x; y ~= x |] ==> x < (y::prat)"; | 
| 5078 | 649 | by (rtac not_prat_leE 1); | 
| 650 | by (fast_tac (claset() addDs [prat_le_imp_less_or_eq]) 1); | |
| 651 | qed "not_less_not_eq_prat_less"; | |
| 652 | ||
| 653 | Goalw [prat_less_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 654 | "[| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::prat)"; | 
| 5078 | 655 | by (REPEAT(etac exE 1)); | 
| 656 | by (res_inst_tac [("x","T+Ta")] exI 1);
 | |
| 657 | by (auto_tac (claset(),simpset() addsimps prat_add_ac)); | |
| 658 | qed "prat_add_less_mono"; | |
| 659 | ||
| 660 | Goalw [prat_less_def] | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 661 | "[| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::prat)"; | 
| 5078 | 662 | by (REPEAT(etac exE 1)); | 
| 663 | by (res_inst_tac [("x","T*Ta+T*x2+x1*Ta")] exI 1);
 | |
| 664 | by (auto_tac (claset(),simpset() addsimps prat_add_ac @ | |
| 665 | [prat_add_mult_distrib,prat_add_mult_distrib2])); | |
| 666 | qed "prat_mult_less_mono"; | |
| 667 | ||
| 668 | (* more prat_le *) | |
| 669 | Goal "!!(q1::prat). q1 <= q2 ==> x * q1 <= x * q2"; | |
| 670 | by (dtac prat_le_imp_less_or_eq 1); | |
| 671 | by (Step_tac 1); | |
| 672 | by (auto_tac (claset() addSIs [prat_le_refl, | |
| 673 | prat_less_imp_le,prat_mult_left_less2_mono1],simpset())); | |
| 674 | qed "prat_mult_left_le2_mono1"; | |
| 675 | ||
| 676 | Goal "!!(q1::prat). q1 <= q2 ==> q1 * x <= q2 * x"; | |
| 677 | by (auto_tac (claset() addDs [prat_mult_left_le2_mono1], | |
| 678 | simpset() addsimps [prat_mult_commute])); | |
| 679 | qed "prat_mult_le2_mono1"; | |
| 680 | ||
| 681 | Goal | |
| 682 | "!!(q1::prat). q1 <= q2 ==> qinv (q2) <= qinv (q1)"; | |
| 683 | by (dtac prat_le_imp_less_or_eq 1); | |
| 684 | by (Step_tac 1); | |
| 685 | by (auto_tac (claset() addSIs [prat_le_refl, | |
| 686 | prat_less_imp_le,qinv_prat_less],simpset())); | |
| 687 | qed "qinv_prat_le"; | |
| 688 | ||
| 689 | Goal "!!(q1::prat). q1 <= q2 ==> x + q1 <= x + q2"; | |
| 690 | by (dtac prat_le_imp_less_or_eq 1); | |
| 691 | by (Step_tac 1); | |
| 692 | by (auto_tac (claset() addSIs [prat_le_refl, | |
| 693 | prat_less_imp_le,prat_add_less2_mono1], | |
| 694 | simpset() addsimps [prat_add_commute])); | |
| 695 | qed "prat_add_left_le2_mono1"; | |
| 696 | ||
| 697 | Goal "!!(q1::prat). q1 <= q2 ==> q1 + x <= q2 + x"; | |
| 698 | by (auto_tac (claset() addDs [prat_add_left_le2_mono1], | |
| 699 | simpset() addsimps [prat_add_commute])); | |
| 700 | qed "prat_add_le2_mono1"; | |
| 701 | ||
| 702 | Goal "!!k l::prat. [|i<=j; k<=l |] ==> i + k <= j + l"; | |
| 703 | by (etac (prat_add_le2_mono1 RS prat_le_trans) 1); | |
| 704 | by (simp_tac (simpset() addsimps [prat_add_commute]) 1); | |
| 705 | (*j moves to the end because it is free while k, l are bound*) | |
| 706 | by (etac prat_add_le2_mono1 1); | |
| 707 | qed "prat_add_le_mono"; | |
| 708 | ||
| 709 | Goal "!!(x::prat). x + y < z + y ==> x < z"; | |
| 710 | by (rtac ccontr 1); | |
| 711 | by (etac (prat_leI RS prat_le_imp_less_or_eq RS disjE) 1); | |
| 712 | by (dres_inst_tac [("x","y"),("q1.0","z")] prat_add_less2_mono1 1);
 | |
| 713 | by (auto_tac (claset() addIs [prat_less_asym], | |
| 714 | simpset() addsimps [prat_less_not_refl])); | |
| 715 | qed "prat_add_right_less_cancel"; | |
| 716 | ||
| 717 | Goal "!!(x::prat). y + x < y + z ==> x < z"; | |
| 718 | by (res_inst_tac [("y","y")] prat_add_right_less_cancel 1);
 | |
| 719 | by (asm_full_simp_tac (simpset() addsimps [prat_add_commute]) 1); | |
| 720 | qed "prat_add_left_less_cancel"; | |
| 721 | ||
| 722 | (*** lemmas required for lemma_gleason9_34 in PReal : w*y > y/z ***) | |
| 723 | Goalw [prat_pnat_def] "Abs_prat(ratrel^^{(x,y)}) = $#x*qinv($#y)";
 | |
| 724 | by (auto_tac (claset(),simpset() addsimps [prat_mult,qinv,pnat_mult_1_left, | |
| 725 | pnat_mult_1])); | |
| 726 | qed "Abs_prat_mult_qinv"; | |
| 727 | ||
| 728 | Goal "Abs_prat(ratrel^^{(x,y)}) <= Abs_prat(ratrel^^{(x,Abs_pnat 1)})";
 | |
| 729 | by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1); | |
| 730 | by (rtac prat_mult_left_le2_mono1 1); | |
| 731 | by (rtac qinv_prat_le 1); | |
| 732 | by (pnat_ind_tac "y" 1); | |
| 733 | by (dres_inst_tac [("x","$#Abs_pnat 1")] prat_add_le2_mono1 2);
 | |
| 734 | by (cut_facts_tac [prat_less_1_2 RS prat_less_imp_le] 2); | |
| 735 | by (auto_tac (claset() addIs [prat_le_trans], | |
| 736 | simpset() addsimps [prat_le_refl, | |
| 737 | pSuc_is_plus_one,pnat_one_def,prat_pnat_add])); | |
| 738 | qed "lemma_Abs_prat_le1"; | |
| 739 | ||
| 740 | Goal "Abs_prat(ratrel^^{(x,Abs_pnat 1)}) <= Abs_prat(ratrel^^{(x*y,Abs_pnat 1)})";
 | |
| 741 | by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1); | |
| 742 | by (rtac prat_mult_le2_mono1 1); | |
| 743 | by (pnat_ind_tac "y" 1); | |
| 744 | by (dres_inst_tac [("x","$#x")] prat_add_le2_mono1 2);
 | |
| 745 | by (cut_inst_tac [("z","$#x")] (prat_self_less_add_self 
 | |
| 746 | RS prat_less_imp_le) 2); | |
| 747 | by (auto_tac (claset() addIs [prat_le_trans], | |
| 748 | simpset() addsimps [prat_le_refl, | |
| 749 | pSuc_is_plus_one,pnat_one_def,prat_add_mult_distrib2, | |
| 750 | prat_pnat_add,prat_pnat_mult])); | |
| 751 | qed "lemma_Abs_prat_le2"; | |
| 752 | ||
| 753 | Goal "Abs_prat(ratrel^^{(x,z)}) <= Abs_prat(ratrel^^{(x*y,Abs_pnat 1)})";
 | |
| 754 | by (fast_tac (claset() addIs [prat_le_trans,lemma_Abs_prat_le1,lemma_Abs_prat_le2]) 1); | |
| 755 | qed "lemma_Abs_prat_le3"; | |
| 756 | ||
| 757 | Goal "Abs_prat(ratrel^^{(x*y,Abs_pnat 1)}) * Abs_prat(ratrel^^{(w,x)}) = \
 | |
| 758 | \         Abs_prat(ratrel^^{(w*y,Abs_pnat 1)})";
 | |
| 759 | by (full_simp_tac (simpset() addsimps [prat_mult, | |
| 760 | pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac) 1); | |
| 761 | qed "pre_lemma_gleason9_34"; | |
| 762 | ||
| 763 | Goal "Abs_prat(ratrel^^{(y*x,Abs_pnat 1*y)}) = \
 | |
| 764 | \         Abs_prat(ratrel^^{(x,Abs_pnat 1)})";
 | |
| 765 | by (auto_tac (claset(),simpset() addsimps | |
| 766 | [pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac)); | |
| 767 | qed "pre_lemma_gleason9_34b"; | |
| 768 | ||
| 769 | Goal "($#n < $#m) = (n < m)"; | |
| 770 | by (auto_tac (claset(),simpset() addsimps [prat_less_def, | |
| 771 | pnat_less_iff,prat_pnat_add])); | |
| 772 | by (res_inst_tac [("z","T")] eq_Abs_prat 1);
 | |
| 773 | by (auto_tac (claset() addDs [pnat_eq_lessI], | |
| 774 | simpset() addsimps [prat_add,pnat_mult_1, | |
| 775 | pnat_mult_1_left,prat_pnat_def,pnat_less_iff RS sym])); | |
| 776 | qed "prat_pnat_less_iff"; | |
| 777 | ||
| 778 | Addsimps [prat_pnat_less_iff]; | |
| 779 | ||
| 780 | (***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***) | |
| 781 | ||
| 782 | (*** prove witness that will be required to prove non-emptiness ***) | |
| 783 | (*** of preal type as defined using Dedekind Sections in PReal ***) | |
| 784 | (*** Show that exists positive real `one' ***) | |
| 785 | ||
| 786 | Goal "? q. q: {x::prat. x < $#Abs_pnat 1}";
 | |
| 787 | by (fast_tac (claset() addIs [prat_less_qinv_2_1]) 1); | |
| 788 | qed "lemma_prat_less_1_memEx"; | |
| 789 | ||
| 790 | Goal "{x::prat. x < $#Abs_pnat 1} ~= {}";
 | |
| 791 | by (rtac notI 1); | |
| 792 | by (cut_facts_tac [lemma_prat_less_1_memEx] 1); | |
| 793 | by (Asm_full_simp_tac 1); | |
| 794 | qed "lemma_prat_less_1_set_non_empty"; | |
| 795 | ||
| 796 | Goalw [psubset_def] "{} < {x::prat. x < $#Abs_pnat 1}";
 | |
| 797 | by (asm_full_simp_tac (simpset() addsimps | |
| 798 | [lemma_prat_less_1_set_non_empty RS not_sym]) 1); | |
| 799 | qed "empty_set_psubset_lemma_prat_less_1_set"; | |
| 800 | ||
| 801 | (*** exists rational not in set --- $#Abs_pnat 1 itself ***) | |
| 802 | Goal "? q. q ~: {x::prat. x < $#Abs_pnat 1}";
 | |
| 803 | by (res_inst_tac [("x","$#Abs_pnat 1")] exI 1);
 | |
| 804 | by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl])); | |
| 805 | qed "lemma_prat_less_1_not_memEx"; | |
| 806 | ||
| 807 | Goal "{x::prat. x < $#Abs_pnat 1} ~= {q::prat. True}";
 | |
| 808 | by (rtac notI 1); | |
| 809 | by (cut_facts_tac [lemma_prat_less_1_not_memEx] 1); | |
| 810 | by (Asm_full_simp_tac 1); | |
| 811 | qed "lemma_prat_less_1_set_not_rat_set"; | |
| 812 | ||
| 813 | Goalw [psubset_def,subset_def] | |
| 814 |       "{x::prat. x < $#Abs_pnat 1} < {q::prat. True}";
 | |
| 815 | by (asm_full_simp_tac (simpset() addsimps | |
| 816 | [lemma_prat_less_1_set_not_rat_set, | |
| 817 | lemma_prat_less_1_not_memEx]) 1); | |
| 818 | qed "lemma_prat_less_1_set_psubset_rat_set"; | |
| 819 | ||
| 820 | (*** prove non_emptiness of type ***) | |
| 821 | Goal "{x::prat. x < $#Abs_pnat 1} : {A. {} < A & A < {q::prat. True} & \
 | |
| 822 | \ (!y: A. ((!z. z < y --> z: A) & \ | |
| 823 | \ (? u: A. y < u)))}"; | |
| 824 | by (auto_tac (claset() addDs [prat_less_trans], | |
| 825 | simpset() addsimps [empty_set_psubset_lemma_prat_less_1_set, | |
| 826 | lemma_prat_less_1_set_psubset_rat_set])); | |
| 827 | by (dtac prat_dense 1); | |
| 828 | by (Fast_tac 1); | |
| 829 | qed "preal_1"; | |
| 830 | ||
| 831 | ||
| 832 | ||
| 833 | ||
| 834 | ||
| 835 |