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