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
|
|
15 |
"!! x1. [| (x1::pnat) * y2 = x2 * y1; x2 * y3 = x3 * y2 |] \
|
|
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]
|
|
33 |
"!!x1 x2. [| x1 * y2 = x2 * y1 |] ==> ((x1,y1),(x2,y2)): ratrel";
|
|
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 |
|
|
146 |
Goal
|
|
147 |
"!!(x1::pnat). [| x1 * y2 = x2 * y1 |] ==> \
|
|
148 |
\ (x * y1 + x1 * ya) * (ya * y2) = (x * y2 + x2 * ya) * (ya * y1)";
|
|
149 |
by (auto_tac (claset() addSIs [pnat_same_multI2],
|
|
150 |
simpset() addsimps [pnat_add_mult_distrib,
|
|
151 |
pnat_mult_assoc]));
|
|
152 |
by (res_inst_tac [("n1","y2")] (pnat_mult_commute RS subst) 1);
|
|
153 |
by (auto_tac (claset() addIs [pnat_add_left_cancel RS iffD2],simpset() addsimps pnat_mult_ac));
|
|
154 |
by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS subst) 1);
|
|
155 |
by (res_inst_tac [("y1","x1")] (pnat_mult_left_commute RS ssubst) 1);
|
|
156 |
by (auto_tac (claset(),simpset() addsimps [pnat_mult_assoc RS sym]));
|
|
157 |
qed "prat_add_congruent2_lemma";
|
|
158 |
|
|
159 |
Goal
|
|
160 |
"congruent2 ratrel (%p1 p2. \
|
|
161 |
\ split (%x1 y1. split (%x2 y2. ratrel^^{(x1*y2 + x2*y1, y1*y2)}) p2) p1)";
|
|
162 |
by (rtac (equiv_ratrel RS congruent2_commuteI) 1);
|
|
163 |
by Safe_tac;
|
|
164 |
by (rewtac split_def);
|
|
165 |
by (asm_simp_tac (simpset() addsimps [pnat_mult_commute,pnat_add_commute]) 1);
|
|
166 |
by (auto_tac (claset(),simpset() addsimps [prat_add_congruent2_lemma]));
|
|
167 |
qed "prat_add_congruent2";
|
|
168 |
|
|
169 |
(*Resolve th against the corresponding facts for prat_add*)
|
|
170 |
val prat_add_ize = RSLIST [equiv_ratrel, prat_add_congruent2];
|
|
171 |
|
|
172 |
Goalw [prat_add_def]
|
|
173 |
"Abs_prat((ratrel^^{(x1,y1)})) + Abs_prat((ratrel^^{(x2,y2)})) = \
|
|
174 |
\ Abs_prat(ratrel ^^ {(x1*y2 + x2*y1, y1*y2)})";
|
|
175 |
by (simp_tac (simpset() addsimps [prat_add_ize UN_equiv_class2]) 1);
|
|
176 |
qed "prat_add";
|
|
177 |
|
|
178 |
Goal "(z::prat) + w = w + z";
|
|
179 |
by (res_inst_tac [("z","z")] eq_Abs_prat 1);
|
|
180 |
by (res_inst_tac [("z","w")] eq_Abs_prat 1);
|
|
181 |
by (asm_simp_tac (simpset() addsimps ([prat_add] @ pnat_add_ac @ pnat_mult_ac)) 1);
|
|
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);
|
|
188 |
by (asm_simp_tac (simpset() addsimps ([pnat_add_mult_distrib2,prat_add] @
|
|
189 |
pnat_add_ac @ pnat_mult_ac)) 1);
|
|
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);
|
|
229 |
by (asm_simp_tac (simpset() addsimps (pnat_mult_ac @ [prat_mult])) 1);
|
|
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 |
|
|
305 |
Goal "!!y. x * y = $# Abs_pnat 1 ==> x = qinv y";
|
|
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
|
|
337 |
(simpset() addsimps ([pnat_add_mult_distrib2, prat_add, prat_mult] @
|
|
338 |
pnat_add_ac @ pnat_mult_ac)) 1);
|
|
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 |
|
|
370 |
Goal
|
|
371 |
"!! Q1. [| Q1 < (Q2::prat); \
|
|
372 |
\ !! (Q3::prat). Q1 + Q3 = Q2 ==> P |] \
|
|
373 |
\ ==> P";
|
|
374 |
by (dtac (prat_lessE_lemma RS mp) 1);
|
|
375 |
by Auto_tac;
|
|
376 |
qed "prat_lessE";
|
|
377 |
|
|
378 |
(* qless is a strong order i.e nonreflexive and transitive *)
|
|
379 |
Goal
|
|
380 |
"!!(q1::prat). [| q1 < q2; q2 < q3 |] ==> q1 < q3";
|
|
381 |
by (REPEAT(dtac (prat_lessE_lemma RS mp) 1));
|
|
382 |
by (REPEAT(etac exE 1));
|
|
383 |
by (hyp_subst_tac 1);
|
|
384 |
by (res_inst_tac [("Q3.0","Q3 + Q3a")] prat_lessI 1);
|
|
385 |
by (auto_tac (claset(),simpset() addsimps [prat_add_assoc]));
|
|
386 |
qed "prat_less_trans";
|
|
387 |
|
|
388 |
Goal "~q < (q::prat)";
|
|
389 |
by (EVERY1[rtac notI, dtac (prat_lessE_lemma RS mp)]);
|
|
390 |
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
|
|
391 |
by (res_inst_tac [("z","Q3")] eq_Abs_prat 1);
|
|
392 |
by (etac exE 1 THEN res_inst_tac [("z","Q3a")] eq_Abs_prat 1);
|
|
393 |
by (REPEAT(hyp_subst_tac 1));
|
|
394 |
by (asm_full_simp_tac (simpset() addsimps [prat_add,
|
|
395 |
pnat_no_add_ident,pnat_add_mult_distrib2] @ pnat_mult_ac) 1);
|
|
396 |
qed "prat_less_not_refl";
|
|
397 |
|
|
398 |
(*** y < y ==> P ***)
|
|
399 |
bind_thm("prat_less_irrefl",prat_less_not_refl RS notE);
|
|
400 |
|
|
401 |
Goal "!! (q1::prat). [| q1 < q2; q2 < q1 |] ==> P";
|
|
402 |
by (dtac prat_less_trans 1 THEN assume_tac 1);
|
|
403 |
by (asm_full_simp_tac (simpset() addsimps [prat_less_not_refl]) 1);
|
|
404 |
qed "prat_less_asym";
|
|
405 |
|
|
406 |
Goal "!! (q1::prat). q1 < q2 ==> ~ q2 < q1";
|
|
407 |
by (auto_tac (claset() addSDs [prat_less_asym],simpset()));
|
|
408 |
qed "prat_less_not_sym";
|
|
409 |
|
|
410 |
(* half of positive fraction exists- Gleason p. 120- Proposition 9-2.6(i)*)
|
|
411 |
Goal "!(q::prat). ? x. x + x = q";
|
|
412 |
by (rtac allI 1);
|
|
413 |
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
|
|
414 |
by (res_inst_tac [("x","Abs_prat (ratrel ^^ {(x, y+y)})")] exI 1);
|
|
415 |
by (auto_tac (claset(),simpset() addsimps
|
|
416 |
[prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
|
|
417 |
pnat_add_mult_distrib2]));
|
|
418 |
qed "lemma_prat_dense";
|
|
419 |
|
|
420 |
Goal "? (x::prat). x + x = q";
|
|
421 |
by (res_inst_tac [("z","q")] eq_Abs_prat 1);
|
|
422 |
by (res_inst_tac [("x","Abs_prat (ratrel ^^ {(x, y+y)})")] exI 1);
|
|
423 |
by (auto_tac (claset(),simpset() addsimps
|
|
424 |
[prat_add,pnat_mult_assoc RS sym,pnat_add_mult_distrib,
|
|
425 |
pnat_add_mult_distrib2]));
|
|
426 |
qed "prat_lemma_dense";
|
|
427 |
|
|
428 |
(* there exists a number between any two positive fractions *)
|
|
429 |
(* Gleason p. 120- Proposition 9-2.6(iv) *)
|
|
430 |
Goalw [prat_less_def]
|
|
431 |
"!! (q1::prat). q1 < q2 ==> ? x. q1 < x & x < q2";
|
|
432 |
by (auto_tac (claset() addIs [lemma_prat_dense],simpset()));
|
|
433 |
by (res_inst_tac [("x","T")] (lemma_prat_dense RS allE) 1);
|
|
434 |
by (etac exE 1);
|
|
435 |
by (res_inst_tac [("x","q1 + x")] exI 1);
|
|
436 |
by (auto_tac (claset() addIs [prat_lemma_dense],
|
|
437 |
simpset() addsimps [prat_add_assoc]));
|
|
438 |
qed "prat_dense";
|
|
439 |
|
|
440 |
(* ordering of addition for positive fractions *)
|
|
441 |
Goalw [prat_less_def]
|
|
442 |
"!!(q1::prat). q1 < q2 ==> q1 + x < q2 + x";
|
|
443 |
by (Step_tac 1);
|
|
444 |
by (res_inst_tac [("x","T")] exI 1);
|
|
445 |
by (auto_tac (claset(),simpset() addsimps prat_add_ac));
|
|
446 |
qed "prat_add_less2_mono1";
|
|
447 |
|
|
448 |
Goal
|
|
449 |
"!!(q1::prat). q1 < q2 ==> x + q1 < x + q2";
|
|
450 |
by (auto_tac (claset() addIs [prat_add_less2_mono1],
|
|
451 |
simpset() addsimps [prat_add_commute]));
|
|
452 |
qed "prat_add_less2_mono2";
|
|
453 |
|
|
454 |
(* ordering of multiplication for positive fractions *)
|
|
455 |
Goalw [prat_less_def]
|
|
456 |
"!!(q1::prat). q1 < q2 ==> q1 * x < q2 * x";
|
|
457 |
by (Step_tac 1);
|
|
458 |
by (res_inst_tac [("x","T*x")] exI 1);
|
|
459 |
by (auto_tac (claset(),simpset() addsimps [prat_add_mult_distrib]));
|
|
460 |
qed "prat_mult_less2_mono1";
|
|
461 |
|
|
462 |
Goal "!!(q1::prat). q1 < q2 ==> x * q1 < x * q2";
|
|
463 |
by (auto_tac (claset() addDs [prat_mult_less2_mono1],
|
|
464 |
simpset() addsimps [prat_mult_commute]));
|
|
465 |
qed "prat_mult_left_less2_mono1";
|
|
466 |
|
|
467 |
(* there is no smallest positive fraction *)
|
|
468 |
Goalw [prat_less_def] "? (x::prat). x < y";
|
|
469 |
by (cut_facts_tac [lemma_prat_dense] 1);
|
|
470 |
by (Fast_tac 1);
|
|
471 |
qed "qless_Ex";
|
|
472 |
|
|
473 |
(* lemma for proving $< is linear *)
|
|
474 |
Goalw [prat_def,prat_less_def]
|
|
475 |
"ratrel ^^ {(x, y * ya)} : {p::(pnat*pnat).True}/ratrel";
|
|
476 |
by (asm_full_simp_tac (simpset() addsimps [ratrel_def,quotient_def]) 1);
|
|
477 |
by (Blast_tac 1);
|
|
478 |
qed "lemma_prat_less_linear";
|
|
479 |
|
|
480 |
(* linearity of < -- Gleason p. 120 - Proposition 9-2.6 *)
|
|
481 |
(*** FIXME Proof long ***)
|
|
482 |
Goalw [prat_less_def]
|
|
483 |
"(q1::prat) < q2 | q1 = q2 | q2 < q1";
|
|
484 |
by (res_inst_tac [("z","q1")] eq_Abs_prat 1);
|
|
485 |
by (res_inst_tac [("z","q2")] eq_Abs_prat 1);
|
|
486 |
by (Step_tac 1 THEN REPEAT(dtac (not_ex RS iffD1) 1)
|
|
487 |
THEN Auto_tac);
|
|
488 |
by (cut_inst_tac [("z1.0","x*ya"),
|
|
489 |
("z2.0","xa*y")] pnat_linear_Ex_eq 1);
|
|
490 |
by (EVERY1[etac disjE,etac exE]);
|
|
491 |
by (eres_inst_tac
|
|
492 |
[("x","Abs_prat(ratrel^^{(xb,ya*y)})")] allE 1);
|
|
493 |
by (asm_full_simp_tac
|
|
494 |
(simpset() addsimps [prat_add, pnat_mult_assoc
|
|
495 |
RS sym,pnat_add_mult_distrib RS sym]) 1);
|
|
496 |
by (EVERY1[asm_full_simp_tac (simpset() addsimps pnat_mult_ac),
|
|
497 |
etac disjE, assume_tac, etac exE]);
|
|
498 |
by (thin_tac "!T. Abs_prat (ratrel ^^ {(x, y)}) + T ~= \
|
|
499 |
\ Abs_prat (ratrel ^^ {(xa, ya)})" 1);
|
|
500 |
by (eres_inst_tac [("x","Abs_prat(ratrel^^{(xb,y*ya)})")] allE 1);
|
|
501 |
by (asm_full_simp_tac (simpset() addsimps [prat_add,
|
|
502 |
pnat_mult_assoc RS sym,pnat_add_mult_distrib RS sym]) 1);
|
|
503 |
by (asm_full_simp_tac (simpset() addsimps pnat_mult_ac) 1);
|
|
504 |
qed "prat_linear";
|
|
505 |
|
|
506 |
Goal
|
|
507 |
"!!(q1::prat). [| q1 < q2 ==> P; q1 = q2 ==> P; \
|
|
508 |
\ q2 < q1 ==> P |] ==> P";
|
|
509 |
by (cut_inst_tac [("q1.0","q1"),("q2.0","q2")] prat_linear 1);
|
|
510 |
by Auto_tac;
|
|
511 |
qed "prat_linear_less2";
|
|
512 |
|
|
513 |
(* Gleason p. 120 -- 9-2.6 (iv) *)
|
|
514 |
Goal
|
|
515 |
"!!(q1::prat). [| q1 < q2; qinv(q1) = qinv(q2) |] ==> P";
|
|
516 |
by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"),
|
|
517 |
("q2.0","q2")] prat_mult_less2_mono1 1);
|
|
518 |
by (assume_tac 1);
|
|
519 |
by (Asm_full_simp_tac 1 THEN dtac sym 1);
|
|
520 |
by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
|
|
521 |
qed "lemma1_qinv_prat_less";
|
|
522 |
|
|
523 |
Goal
|
|
524 |
"!!(q1::prat). [| q1 < q2; qinv(q1) < qinv(q2) |] ==> P";
|
|
525 |
by (cut_inst_tac [("x","qinv (q2)"),("q1.0","q1"),
|
|
526 |
("q2.0","q2")] prat_mult_less2_mono1 1);
|
|
527 |
by (assume_tac 1);
|
|
528 |
by (cut_inst_tac [("x","q1"),("q1.0","qinv (q1)"),
|
|
529 |
("q2.0","qinv (q2)")] prat_mult_left_less2_mono1 1);
|
|
530 |
by Auto_tac;
|
|
531 |
by (dres_inst_tac [("q2.0","$#Abs_pnat 1")] prat_less_trans 1);
|
|
532 |
by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
|
|
533 |
qed "lemma2_qinv_prat_less";
|
|
534 |
|
|
535 |
Goal
|
|
536 |
"!!(q1::prat). q1 < q2 ==> qinv (q2) < qinv (q1)";
|
|
537 |
by (res_inst_tac [("q2.0","qinv q1"),
|
|
538 |
("q1.0","qinv q2")] prat_linear_less2 1);
|
|
539 |
by (auto_tac (claset() addEs [lemma1_qinv_prat_less,
|
|
540 |
lemma2_qinv_prat_less],simpset()));
|
|
541 |
qed "qinv_prat_less";
|
|
542 |
|
|
543 |
Goal "!!(q1::prat). q1 < $#Abs_pnat 1 ==> $#Abs_pnat 1 < qinv(q1)";
|
|
544 |
by (dtac qinv_prat_less 1);
|
|
545 |
by (full_simp_tac (simpset() addsimps [qinv_1]) 1);
|
|
546 |
qed "prat_qinv_gt_1";
|
|
547 |
|
|
548 |
Goalw [pnat_one_def] "!!(q1::prat). q1 < $#1p ==> $#1p < qinv(q1)";
|
|
549 |
by (etac prat_qinv_gt_1 1);
|
|
550 |
qed "prat_qinv_is_gt_1";
|
|
551 |
|
|
552 |
Goalw [prat_less_def] "$#Abs_pnat 1 < $#Abs_pnat 1 + $#Abs_pnat 1";
|
|
553 |
by (Fast_tac 1);
|
|
554 |
qed "prat_less_1_2";
|
|
555 |
|
|
556 |
Goal "qinv($#Abs_pnat 1 + $#Abs_pnat 1) < $#Abs_pnat 1";
|
|
557 |
by (cut_facts_tac [prat_less_1_2 RS qinv_prat_less] 1);
|
|
558 |
by (asm_full_simp_tac (simpset() addsimps [qinv_1]) 1);
|
|
559 |
qed "prat_less_qinv_2_1";
|
|
560 |
|
|
561 |
Goal "!!(x::prat). x < y ==> x*qinv(y) < $#Abs_pnat 1";
|
|
562 |
by (dres_inst_tac [("x","qinv(y)")] prat_mult_less2_mono1 1);
|
|
563 |
by (Asm_full_simp_tac 1);
|
|
564 |
qed "prat_mult_qinv_less_1";
|
|
565 |
|
|
566 |
Goal "(x::prat) < x + x";
|
|
567 |
by (cut_inst_tac [("x","x")]
|
|
568 |
(prat_less_1_2 RS prat_mult_left_less2_mono1) 1);
|
|
569 |
by (asm_full_simp_tac (simpset() addsimps
|
|
570 |
[prat_add_mult_distrib2]) 1);
|
|
571 |
qed "prat_self_less_add_self";
|
|
572 |
|
|
573 |
Goalw [prat_less_def] "(x::prat) < y + x";
|
|
574 |
by (res_inst_tac [("x","y")] exI 1);
|
|
575 |
by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
|
|
576 |
qed "prat_self_less_add_right";
|
|
577 |
|
|
578 |
Goal "(x::prat) < x + y";
|
|
579 |
by (rtac (prat_add_commute RS subst) 1);
|
|
580 |
by (simp_tac (simpset() addsimps [prat_self_less_add_right]) 1);
|
|
581 |
qed "prat_self_less_add_left";
|
|
582 |
|
|
583 |
Goalw [prat_less_def] "!!y. $#1p < y ==> (x::prat) < x * y";
|
|
584 |
by (auto_tac (claset(),simpset() addsimps [pnat_one_def,
|
|
585 |
prat_add_mult_distrib2]));
|
|
586 |
qed "prat_self_less_mult_right";
|
|
587 |
|
|
588 |
(*** Properties of <= ***)
|
|
589 |
|
|
590 |
Goalw [prat_le_def] "!!w. ~(w < z) ==> z <= (w::prat)";
|
|
591 |
by (assume_tac 1);
|
|
592 |
qed "prat_leI";
|
|
593 |
|
|
594 |
Goalw [prat_le_def] "!!w. z<=w ==> ~(w<(z::prat))";
|
|
595 |
by (assume_tac 1);
|
|
596 |
qed "prat_leD";
|
|
597 |
|
|
598 |
val prat_leE = make_elim prat_leD;
|
|
599 |
|
|
600 |
Goal "!!w. (~(w < z)) = (z <= (w::prat))";
|
|
601 |
by (fast_tac (claset() addSIs [prat_leI,prat_leD]) 1);
|
|
602 |
qed "prat_less_le_iff";
|
|
603 |
|
|
604 |
Goalw [prat_le_def] "!!z. ~ z <= w ==> w<(z::prat)";
|
|
605 |
by (Fast_tac 1);
|
|
606 |
qed "not_prat_leE";
|
|
607 |
|
|
608 |
Goalw [prat_le_def] "!!z. z < w ==> z <= (w::prat)";
|
|
609 |
by (fast_tac (claset() addEs [prat_less_asym]) 1);
|
|
610 |
qed "prat_less_imp_le";
|
|
611 |
|
|
612 |
Goalw [prat_le_def] "!!(x::prat). x <= y ==> x < y | x = y";
|
|
613 |
by (cut_facts_tac [prat_linear] 1);
|
|
614 |
by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
|
|
615 |
qed "prat_le_imp_less_or_eq";
|
|
616 |
|
|
617 |
Goalw [prat_le_def] "!!z. z<w | z=w ==> z <=(w::prat)";
|
|
618 |
by (cut_facts_tac [prat_linear] 1);
|
|
619 |
by (fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym]) 1);
|
|
620 |
qed "prat_less_or_eq_imp_le";
|
|
621 |
|
|
622 |
Goal "(x <= (y::prat)) = (x < y | x=y)";
|
|
623 |
by (REPEAT(ares_tac [iffI, prat_less_or_eq_imp_le, prat_le_imp_less_or_eq] 1));
|
|
624 |
qed "prat_le_eq_less_or_eq";
|
|
625 |
|
|
626 |
Goal "w <= (w::prat)";
|
|
627 |
by (simp_tac (simpset() addsimps [prat_le_eq_less_or_eq]) 1);
|
|
628 |
qed "prat_le_refl";
|
|
629 |
|
|
630 |
val prems = goal thy "!!i. [| i <= j; j < k |] ==> i < (k::prat)";
|
|
631 |
by (dtac prat_le_imp_less_or_eq 1);
|
|
632 |
by (fast_tac (claset() addIs [prat_less_trans]) 1);
|
|
633 |
qed "prat_le_less_trans";
|
|
634 |
|
|
635 |
Goal "!! (i::prat). [| i < j; j <= k |] ==> i < k";
|
|
636 |
by (dtac prat_le_imp_less_or_eq 1);
|
|
637 |
by (fast_tac (claset() addIs [prat_less_trans]) 1);
|
|
638 |
qed "prat_less_le_trans";
|
|
639 |
|
|
640 |
Goal "!!i. [| i <= j; j <= k |] ==> i <= (k::prat)";
|
|
641 |
by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq,
|
|
642 |
rtac prat_less_or_eq_imp_le, fast_tac (claset() addIs [prat_less_trans])]);
|
|
643 |
qed "prat_le_trans";
|
|
644 |
|
|
645 |
Goal "!!z. [| z <= w; w <= z |] ==> z = (w::prat)";
|
|
646 |
by (EVERY1 [dtac prat_le_imp_less_or_eq, dtac prat_le_imp_less_or_eq,
|
|
647 |
fast_tac (claset() addEs [prat_less_irrefl,prat_less_asym])]);
|
|
648 |
qed "prat_le_anti_sym";
|
|
649 |
|
|
650 |
Goal "!!x. [| ~ y < x; y ~= x |] ==> x < (y::prat)";
|
|
651 |
by (rtac not_prat_leE 1);
|
|
652 |
by (fast_tac (claset() addDs [prat_le_imp_less_or_eq]) 1);
|
|
653 |
qed "not_less_not_eq_prat_less";
|
|
654 |
|
|
655 |
Goalw [prat_less_def]
|
|
656 |
"!!x. [| x1 < y1; x2 < y2 |] ==> x1 + x2 < y1 + (y2::prat)";
|
|
657 |
by (REPEAT(etac exE 1));
|
|
658 |
by (res_inst_tac [("x","T+Ta")] exI 1);
|
|
659 |
by (auto_tac (claset(),simpset() addsimps prat_add_ac));
|
|
660 |
qed "prat_add_less_mono";
|
|
661 |
|
|
662 |
Goalw [prat_less_def]
|
|
663 |
"!!x. [| x1 < y1; x2 < y2 |] ==> x1 * x2 < y1 * (y2::prat)";
|
|
664 |
by (REPEAT(etac exE 1));
|
|
665 |
by (res_inst_tac [("x","T*Ta+T*x2+x1*Ta")] exI 1);
|
|
666 |
by (auto_tac (claset(),simpset() addsimps prat_add_ac @
|
|
667 |
[prat_add_mult_distrib,prat_add_mult_distrib2]));
|
|
668 |
qed "prat_mult_less_mono";
|
|
669 |
|
|
670 |
(* more prat_le *)
|
|
671 |
Goal "!!(q1::prat). q1 <= q2 ==> x * q1 <= x * q2";
|
|
672 |
by (dtac prat_le_imp_less_or_eq 1);
|
|
673 |
by (Step_tac 1);
|
|
674 |
by (auto_tac (claset() addSIs [prat_le_refl,
|
|
675 |
prat_less_imp_le,prat_mult_left_less2_mono1],simpset()));
|
|
676 |
qed "prat_mult_left_le2_mono1";
|
|
677 |
|
|
678 |
Goal "!!(q1::prat). q1 <= q2 ==> q1 * x <= q2 * x";
|
|
679 |
by (auto_tac (claset() addDs [prat_mult_left_le2_mono1],
|
|
680 |
simpset() addsimps [prat_mult_commute]));
|
|
681 |
qed "prat_mult_le2_mono1";
|
|
682 |
|
|
683 |
Goal
|
|
684 |
"!!(q1::prat). q1 <= q2 ==> qinv (q2) <= qinv (q1)";
|
|
685 |
by (dtac prat_le_imp_less_or_eq 1);
|
|
686 |
by (Step_tac 1);
|
|
687 |
by (auto_tac (claset() addSIs [prat_le_refl,
|
|
688 |
prat_less_imp_le,qinv_prat_less],simpset()));
|
|
689 |
qed "qinv_prat_le";
|
|
690 |
|
|
691 |
Goal "!!(q1::prat). q1 <= q2 ==> x + q1 <= x + q2";
|
|
692 |
by (dtac prat_le_imp_less_or_eq 1);
|
|
693 |
by (Step_tac 1);
|
|
694 |
by (auto_tac (claset() addSIs [prat_le_refl,
|
|
695 |
prat_less_imp_le,prat_add_less2_mono1],
|
|
696 |
simpset() addsimps [prat_add_commute]));
|
|
697 |
qed "prat_add_left_le2_mono1";
|
|
698 |
|
|
699 |
Goal "!!(q1::prat). q1 <= q2 ==> q1 + x <= q2 + x";
|
|
700 |
by (auto_tac (claset() addDs [prat_add_left_le2_mono1],
|
|
701 |
simpset() addsimps [prat_add_commute]));
|
|
702 |
qed "prat_add_le2_mono1";
|
|
703 |
|
|
704 |
Goal "!!k l::prat. [|i<=j; k<=l |] ==> i + k <= j + l";
|
|
705 |
by (etac (prat_add_le2_mono1 RS prat_le_trans) 1);
|
|
706 |
by (simp_tac (simpset() addsimps [prat_add_commute]) 1);
|
|
707 |
(*j moves to the end because it is free while k, l are bound*)
|
|
708 |
by (etac prat_add_le2_mono1 1);
|
|
709 |
qed "prat_add_le_mono";
|
|
710 |
|
|
711 |
Goal "!!(x::prat). x + y < z + y ==> x < z";
|
|
712 |
by (rtac ccontr 1);
|
|
713 |
by (etac (prat_leI RS prat_le_imp_less_or_eq RS disjE) 1);
|
|
714 |
by (dres_inst_tac [("x","y"),("q1.0","z")] prat_add_less2_mono1 1);
|
|
715 |
by (auto_tac (claset() addIs [prat_less_asym],
|
|
716 |
simpset() addsimps [prat_less_not_refl]));
|
|
717 |
qed "prat_add_right_less_cancel";
|
|
718 |
|
|
719 |
Goal "!!(x::prat). y + x < y + z ==> x < z";
|
|
720 |
by (res_inst_tac [("y","y")] prat_add_right_less_cancel 1);
|
|
721 |
by (asm_full_simp_tac (simpset() addsimps [prat_add_commute]) 1);
|
|
722 |
qed "prat_add_left_less_cancel";
|
|
723 |
|
|
724 |
(*** lemmas required for lemma_gleason9_34 in PReal : w*y > y/z ***)
|
|
725 |
Goalw [prat_pnat_def] "Abs_prat(ratrel^^{(x,y)}) = $#x*qinv($#y)";
|
|
726 |
by (auto_tac (claset(),simpset() addsimps [prat_mult,qinv,pnat_mult_1_left,
|
|
727 |
pnat_mult_1]));
|
|
728 |
qed "Abs_prat_mult_qinv";
|
|
729 |
|
|
730 |
Goal "Abs_prat(ratrel^^{(x,y)}) <= Abs_prat(ratrel^^{(x,Abs_pnat 1)})";
|
|
731 |
by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
|
|
732 |
by (rtac prat_mult_left_le2_mono1 1);
|
|
733 |
by (rtac qinv_prat_le 1);
|
|
734 |
by (pnat_ind_tac "y" 1);
|
|
735 |
by (dres_inst_tac [("x","$#Abs_pnat 1")] prat_add_le2_mono1 2);
|
|
736 |
by (cut_facts_tac [prat_less_1_2 RS prat_less_imp_le] 2);
|
|
737 |
by (auto_tac (claset() addIs [prat_le_trans],
|
|
738 |
simpset() addsimps [prat_le_refl,
|
|
739 |
pSuc_is_plus_one,pnat_one_def,prat_pnat_add]));
|
|
740 |
qed "lemma_Abs_prat_le1";
|
|
741 |
|
|
742 |
Goal "Abs_prat(ratrel^^{(x,Abs_pnat 1)}) <= Abs_prat(ratrel^^{(x*y,Abs_pnat 1)})";
|
|
743 |
by (simp_tac (simpset() addsimps [Abs_prat_mult_qinv]) 1);
|
|
744 |
by (rtac prat_mult_le2_mono1 1);
|
|
745 |
by (pnat_ind_tac "y" 1);
|
|
746 |
by (dres_inst_tac [("x","$#x")] prat_add_le2_mono1 2);
|
|
747 |
by (cut_inst_tac [("z","$#x")] (prat_self_less_add_self
|
|
748 |
RS prat_less_imp_le) 2);
|
|
749 |
by (auto_tac (claset() addIs [prat_le_trans],
|
|
750 |
simpset() addsimps [prat_le_refl,
|
|
751 |
pSuc_is_plus_one,pnat_one_def,prat_add_mult_distrib2,
|
|
752 |
prat_pnat_add,prat_pnat_mult]));
|
|
753 |
qed "lemma_Abs_prat_le2";
|
|
754 |
|
|
755 |
Goal "Abs_prat(ratrel^^{(x,z)}) <= Abs_prat(ratrel^^{(x*y,Abs_pnat 1)})";
|
|
756 |
by (fast_tac (claset() addIs [prat_le_trans,lemma_Abs_prat_le1,lemma_Abs_prat_le2]) 1);
|
|
757 |
qed "lemma_Abs_prat_le3";
|
|
758 |
|
|
759 |
Goal "Abs_prat(ratrel^^{(x*y,Abs_pnat 1)}) * Abs_prat(ratrel^^{(w,x)}) = \
|
|
760 |
\ Abs_prat(ratrel^^{(w*y,Abs_pnat 1)})";
|
|
761 |
by (full_simp_tac (simpset() addsimps [prat_mult,
|
|
762 |
pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac) 1);
|
|
763 |
qed "pre_lemma_gleason9_34";
|
|
764 |
|
|
765 |
Goal "Abs_prat(ratrel^^{(y*x,Abs_pnat 1*y)}) = \
|
|
766 |
\ Abs_prat(ratrel^^{(x,Abs_pnat 1)})";
|
|
767 |
by (auto_tac (claset(),simpset() addsimps
|
|
768 |
[pnat_mult_1,pnat_mult_1_left] @ pnat_mult_ac));
|
|
769 |
qed "pre_lemma_gleason9_34b";
|
|
770 |
|
|
771 |
Goal "($#n < $#m) = (n < m)";
|
|
772 |
by (auto_tac (claset(),simpset() addsimps [prat_less_def,
|
|
773 |
pnat_less_iff,prat_pnat_add]));
|
|
774 |
by (res_inst_tac [("z","T")] eq_Abs_prat 1);
|
|
775 |
by (auto_tac (claset() addDs [pnat_eq_lessI],
|
|
776 |
simpset() addsimps [prat_add,pnat_mult_1,
|
|
777 |
pnat_mult_1_left,prat_pnat_def,pnat_less_iff RS sym]));
|
|
778 |
qed "prat_pnat_less_iff";
|
|
779 |
|
|
780 |
Addsimps [prat_pnat_less_iff];
|
|
781 |
|
|
782 |
(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)(***)
|
|
783 |
|
|
784 |
(*** prove witness that will be required to prove non-emptiness ***)
|
|
785 |
(*** of preal type as defined using Dedekind Sections in PReal ***)
|
|
786 |
(*** Show that exists positive real `one' ***)
|
|
787 |
|
|
788 |
Goal "? q. q: {x::prat. x < $#Abs_pnat 1}";
|
|
789 |
by (fast_tac (claset() addIs [prat_less_qinv_2_1]) 1);
|
|
790 |
qed "lemma_prat_less_1_memEx";
|
|
791 |
|
|
792 |
Goal "{x::prat. x < $#Abs_pnat 1} ~= {}";
|
|
793 |
by (rtac notI 1);
|
|
794 |
by (cut_facts_tac [lemma_prat_less_1_memEx] 1);
|
|
795 |
by (Asm_full_simp_tac 1);
|
|
796 |
qed "lemma_prat_less_1_set_non_empty";
|
|
797 |
|
|
798 |
Goalw [psubset_def] "{} < {x::prat. x < $#Abs_pnat 1}";
|
|
799 |
by (asm_full_simp_tac (simpset() addsimps
|
|
800 |
[lemma_prat_less_1_set_non_empty RS not_sym]) 1);
|
|
801 |
qed "empty_set_psubset_lemma_prat_less_1_set";
|
|
802 |
|
|
803 |
(*** exists rational not in set --- $#Abs_pnat 1 itself ***)
|
|
804 |
Goal "? q. q ~: {x::prat. x < $#Abs_pnat 1}";
|
|
805 |
by (res_inst_tac [("x","$#Abs_pnat 1")] exI 1);
|
|
806 |
by (auto_tac (claset(),simpset() addsimps [prat_less_not_refl]));
|
|
807 |
qed "lemma_prat_less_1_not_memEx";
|
|
808 |
|
|
809 |
Goal "{x::prat. x < $#Abs_pnat 1} ~= {q::prat. True}";
|
|
810 |
by (rtac notI 1);
|
|
811 |
by (cut_facts_tac [lemma_prat_less_1_not_memEx] 1);
|
|
812 |
by (Asm_full_simp_tac 1);
|
|
813 |
qed "lemma_prat_less_1_set_not_rat_set";
|
|
814 |
|
|
815 |
Goalw [psubset_def,subset_def]
|
|
816 |
"{x::prat. x < $#Abs_pnat 1} < {q::prat. True}";
|
|
817 |
by (asm_full_simp_tac (simpset() addsimps
|
|
818 |
[lemma_prat_less_1_set_not_rat_set,
|
|
819 |
lemma_prat_less_1_not_memEx]) 1);
|
|
820 |
qed "lemma_prat_less_1_set_psubset_rat_set";
|
|
821 |
|
|
822 |
(*** prove non_emptiness of type ***)
|
|
823 |
Goal "{x::prat. x < $#Abs_pnat 1} : {A. {} < A & A < {q::prat. True} & \
|
|
824 |
\ (!y: A. ((!z. z < y --> z: A) & \
|
|
825 |
\ (? u: A. y < u)))}";
|
|
826 |
by (auto_tac (claset() addDs [prat_less_trans],
|
|
827 |
simpset() addsimps [empty_set_psubset_lemma_prat_less_1_set,
|
|
828 |
lemma_prat_less_1_set_psubset_rat_set]));
|
|
829 |
by (dtac prat_dense 1);
|
|
830 |
by (Fast_tac 1);
|
|
831 |
qed "preal_1";
|
|
832 |
|
|
833 |
|
|
834 |
|
|
835 |
|
|
836 |
|
|
837 |
|