437
|
1 |
(* Title: ZF/OrderArith.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1994 University of Cambridge
|
|
5 |
|
|
6 |
Towards ordinal arithmetic
|
|
7 |
*)
|
|
8 |
|
|
9 |
(*for deleting an unwanted assumption*)
|
|
10 |
val thin = prove_goal pure_thy "[| PROP P; PROP Q |] ==> PROP Q"
|
|
11 |
(fn prems => [resolve_tac prems 1]);
|
|
12 |
|
|
13 |
open OrderArith;
|
|
14 |
|
|
15 |
|
|
16 |
(**** Addition of relations -- disjoint sum ****)
|
|
17 |
|
|
18 |
(** Rewrite rules. Can be used to obtain introduction rules **)
|
|
19 |
|
|
20 |
goalw OrderArith.thy [radd_def]
|
|
21 |
"<Inl(a), Inr(b)> : radd(A,r,B,s) <-> a:A & b:B";
|
|
22 |
by (fast_tac sum_cs 1);
|
|
23 |
val radd_Inl_Inr_iff = result();
|
|
24 |
|
|
25 |
goalw OrderArith.thy [radd_def]
|
|
26 |
"<Inl(a'), Inl(a)> : radd(A,r,B,s) <-> <a',a>:r & a':A & a:A";
|
|
27 |
by (fast_tac sum_cs 1);
|
|
28 |
val radd_Inl_iff = result();
|
|
29 |
|
|
30 |
goalw OrderArith.thy [radd_def]
|
|
31 |
"<Inr(b'), Inr(b)> : radd(A,r,B,s) <-> <b',b>:s & b':B & b:B";
|
|
32 |
by (fast_tac sum_cs 1);
|
|
33 |
val radd_Inr_iff = result();
|
|
34 |
|
|
35 |
goalw OrderArith.thy [radd_def]
|
|
36 |
"<Inr(b), Inl(a)> : radd(A,r,B,s) <-> False";
|
|
37 |
by (fast_tac sum_cs 1);
|
|
38 |
val radd_Inr_Inl_iff = result();
|
|
39 |
|
|
40 |
(** Elimination Rule **)
|
|
41 |
|
|
42 |
val major::prems = goalw OrderArith.thy [radd_def]
|
|
43 |
"[| <p',p> : radd(A,r,B,s); \
|
|
44 |
\ !!x y. [| p'=Inl(x); x:A; p=Inr(y); y:B |] ==> Q; \
|
|
45 |
\ !!x' x. [| p'=Inl(x'); p=Inl(x); <x',x>: r; x':A; x:A |] ==> Q; \
|
|
46 |
\ !!y' y. [| p'=Inr(y'); p=Inr(y); <y',y>: s; y':B; y:B |] ==> Q \
|
|
47 |
\ |] ==> Q";
|
|
48 |
by (cut_facts_tac [major] 1);
|
|
49 |
(*Split into the three cases*)
|
|
50 |
by (REPEAT_FIRST
|
|
51 |
(eresolve_tac [CollectE, Pair_inject, conjE, exE, SigmaE, disjE]));
|
|
52 |
(*Apply each premise to correct subgoal; can't just use fast_tac
|
|
53 |
because hyp_subst_tac would delete equalities too quickly*)
|
|
54 |
by (EVERY (map (fn prem =>
|
|
55 |
EVERY1 [rtac prem, assume_tac, REPEAT o fast_tac sum_cs])
|
|
56 |
prems));
|
|
57 |
val raddE = result();
|
|
58 |
|
|
59 |
(** Type checking **)
|
|
60 |
|
|
61 |
goalw OrderArith.thy [radd_def] "radd(A,r,B,s) <= (A+B) * (A+B)";
|
|
62 |
by (rtac Collect_subset 1);
|
|
63 |
val radd_type = result();
|
|
64 |
|
|
65 |
val field_radd = standard (radd_type RS field_rel_subset);
|
|
66 |
|
|
67 |
(** Linearity **)
|
|
68 |
|
|
69 |
val sum_ss = ZF_ss addsimps [Pair_iff, InlI, InrI, Inl_iff, Inr_iff,
|
|
70 |
Inl_Inr_iff, Inr_Inl_iff];
|
|
71 |
|
|
72 |
val radd_ss = sum_ss addsimps [radd_Inl_iff, radd_Inr_iff,
|
|
73 |
radd_Inl_Inr_iff, radd_Inr_Inl_iff];
|
|
74 |
|
|
75 |
goalw OrderArith.thy [linear_def]
|
|
76 |
"!!r s. [| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))";
|
|
77 |
by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac sumE));
|
|
78 |
by (ALLGOALS (asm_simp_tac radd_ss));
|
|
79 |
val linear_radd = result();
|
|
80 |
|
|
81 |
|
|
82 |
(** Well-foundedness **)
|
|
83 |
|
|
84 |
goal OrderArith.thy
|
|
85 |
"!!r s. [| wf[A](r); wf[B](s) |] ==> wf[A+B](radd(A,r,B,s))";
|
|
86 |
by (rtac wf_onI2 1);
|
|
87 |
by (subgoal_tac "ALL x:A. Inl(x): Ba" 1);
|
|
88 |
(*Proving the lemma, which is needed twice!*)
|
|
89 |
by (eres_inst_tac [("P", "y : A + B")] thin 2);
|
|
90 |
by (rtac ballI 2);
|
|
91 |
by (eres_inst_tac [("r","r"),("a","x")] wf_on_induct 2 THEN assume_tac 2);
|
|
92 |
by (etac (bspec RS mp) 2);
|
|
93 |
by (fast_tac sum_cs 2);
|
|
94 |
by (best_tac (sum_cs addSEs [raddE]) 2);
|
|
95 |
(*Returning to main part of proof*)
|
|
96 |
by (REPEAT_FIRST (eresolve_tac [sumE, ssubst]));
|
|
97 |
by (best_tac sum_cs 1);
|
|
98 |
by (eres_inst_tac [("r","s"),("a","ya")] wf_on_induct 1 THEN assume_tac 1);
|
|
99 |
by (etac (bspec RS mp) 1);
|
|
100 |
by (fast_tac sum_cs 1);
|
|
101 |
by (best_tac (sum_cs addSEs [raddE]) 1);
|
|
102 |
val wf_on_radd = result();
|
|
103 |
|
|
104 |
goal OrderArith.thy
|
|
105 |
"!!r s. [| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))";
|
|
106 |
by (asm_full_simp_tac (ZF_ss addsimps [wf_iff_wf_on_field]) 1);
|
|
107 |
by (rtac (field_radd RSN (2, wf_on_subset_A)) 1);
|
|
108 |
by (REPEAT (ares_tac [wf_on_radd] 1));
|
|
109 |
val wf_radd = result();
|
|
110 |
|
|
111 |
goal OrderArith.thy
|
|
112 |
"!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \
|
|
113 |
\ well_ord(A+B, radd(A,r,B,s))";
|
|
114 |
by (rtac well_ordI 1);
|
|
115 |
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, wf_on_radd]) 1);
|
|
116 |
by (asm_full_simp_tac
|
|
117 |
(ZF_ss addsimps [well_ord_def, tot_ord_def, linear_radd]) 1);
|
|
118 |
val well_ord_radd = result();
|
|
119 |
|
|
120 |
|
|
121 |
(**** Multiplication of relations -- lexicographic product ****)
|
|
122 |
|
|
123 |
(** Rewrite rule. Can be used to obtain introduction rules **)
|
|
124 |
|
|
125 |
goalw OrderArith.thy [rmult_def]
|
|
126 |
"!!r s. <<a',b'>, <a,b>> : rmult(A,r,B,s) <-> \
|
|
127 |
\ (<a',a>: r & a':A & a:A & b': B & b: B) | \
|
|
128 |
\ (<b',b>: s & a'=a & a:A & b': B & b: B)";
|
|
129 |
by (fast_tac ZF_cs 1);
|
|
130 |
val rmult_iff = result();
|
|
131 |
|
|
132 |
val major::prems = goal OrderArith.thy
|
|
133 |
"[| <<a',b'>, <a,b>> : rmult(A,r,B,s); \
|
|
134 |
\ [| <a',a>: r; a':A; a:A; b':B; b:B |] ==> Q; \
|
|
135 |
\ [| <b',b>: s; a:A; a'=a; b':B; b:B |] ==> Q \
|
|
136 |
\ |] ==> Q";
|
|
137 |
by (rtac (major RS (rmult_iff RS iffD1) RS disjE) 1);
|
|
138 |
by (DEPTH_SOLVE (eresolve_tac ([asm_rl, conjE] @ prems) 1));
|
|
139 |
val rmultE = result();
|
|
140 |
|
|
141 |
(** Type checking **)
|
|
142 |
|
|
143 |
goalw OrderArith.thy [rmult_def] "rmult(A,r,B,s) <= (A*B) * (A*B)";
|
|
144 |
by (rtac Collect_subset 1);
|
|
145 |
val rmult_type = result();
|
|
146 |
|
|
147 |
val field_rmult = standard (rmult_type RS field_rel_subset);
|
|
148 |
|
|
149 |
(** Linearity **)
|
|
150 |
|
|
151 |
val [lina,linb] = goal OrderArith.thy
|
|
152 |
"[| linear(A,r); linear(B,s) |] ==> linear(A*B,rmult(A,r,B,s))";
|
|
153 |
by (rewtac linear_def); (*Note! the premises are NOT rewritten*)
|
|
154 |
by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac SigmaE));
|
|
155 |
by (asm_simp_tac (ZF_ss addsimps [rmult_iff]) 1);
|
|
156 |
by (res_inst_tac [("x","xa"), ("y","xb")] (lina RS linearE) 1);
|
|
157 |
by (res_inst_tac [("x","ya"), ("y","yb")] (linb RS linearE) 4);
|
|
158 |
by (REPEAT_SOME (fast_tac ZF_cs));
|
|
159 |
val linear_rmult = result();
|
|
160 |
|
|
161 |
|
|
162 |
(** Well-foundedness **)
|
|
163 |
|
|
164 |
goal OrderArith.thy
|
|
165 |
"!!r s. [| wf[A](r); wf[B](s) |] ==> wf[A*B](rmult(A,r,B,s))";
|
|
166 |
by (rtac wf_onI2 1);
|
|
167 |
by (etac SigmaE 1);
|
|
168 |
by (etac ssubst 1);
|
|
169 |
by (subgoal_tac "ALL b:B. <x,b>: Ba" 1);
|
|
170 |
by (fast_tac ZF_cs 1);
|
|
171 |
by (eres_inst_tac [("a","x")] wf_on_induct 1 THEN assume_tac 1);
|
|
172 |
by (rtac ballI 1);
|
|
173 |
by (eres_inst_tac [("a","b")] wf_on_induct 1 THEN assume_tac 1);
|
|
174 |
by (etac (bspec RS mp) 1);
|
|
175 |
by (fast_tac ZF_cs 1);
|
|
176 |
by (best_tac (ZF_cs addSEs [rmultE]) 1);
|
|
177 |
val wf_on_rmult = result();
|
|
178 |
|
|
179 |
|
|
180 |
goal OrderArith.thy
|
|
181 |
"!!r s. [| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))";
|
|
182 |
by (asm_full_simp_tac (ZF_ss addsimps [wf_iff_wf_on_field]) 1);
|
|
183 |
by (rtac (field_rmult RSN (2, wf_on_subset_A)) 1);
|
|
184 |
by (REPEAT (ares_tac [wf_on_rmult] 1));
|
|
185 |
val wf_rmult = result();
|
|
186 |
|
|
187 |
goal OrderArith.thy
|
|
188 |
"!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \
|
|
189 |
\ well_ord(A*B, rmult(A,r,B,s))";
|
|
190 |
by (rtac well_ordI 1);
|
|
191 |
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, wf_on_rmult]) 1);
|
|
192 |
by (asm_full_simp_tac
|
|
193 |
(ZF_ss addsimps [well_ord_def, tot_ord_def, linear_rmult]) 1);
|
|
194 |
val well_ord_rmult = result();
|
|
195 |
|
|
196 |
|
|
197 |
(**** Inverse image of a relation ****)
|
|
198 |
|
|
199 |
(** Rewrite rule **)
|
|
200 |
|
|
201 |
goalw OrderArith.thy [rvimage_def]
|
|
202 |
"<a,b> : rvimage(A,f,r) <-> <f`a,f`b>: r & a:A & b:A";
|
|
203 |
by (fast_tac ZF_cs 1);
|
|
204 |
val rvimage_iff = result();
|
|
205 |
|
|
206 |
(** Type checking **)
|
|
207 |
|
|
208 |
goalw OrderArith.thy [rvimage_def] "rvimage(A,f,r) <= A*A";
|
|
209 |
by (rtac Collect_subset 1);
|
|
210 |
val rvimage_type = result();
|
|
211 |
|
|
212 |
val field_rvimage = standard (rvimage_type RS field_rel_subset);
|
|
213 |
|
|
214 |
|
|
215 |
(** Linearity **)
|
|
216 |
|
|
217 |
val [finj,lin] = goalw OrderArith.thy [inj_def]
|
|
218 |
"[| f: inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))";
|
|
219 |
by (rewtac linear_def); (*Note! the premises are NOT rewritten*)
|
|
220 |
by (REPEAT_FIRST (ares_tac [ballI]));
|
|
221 |
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff]) 1);
|
|
222 |
by (cut_facts_tac [finj] 1);
|
|
223 |
by (res_inst_tac [("x","f`x"), ("y","f`y")] (lin RS linearE) 1);
|
|
224 |
by (REPEAT_SOME (fast_tac (ZF_cs addSIs [apply_type])));
|
|
225 |
val linear_rvimage = result();
|
|
226 |
|
|
227 |
|
|
228 |
(** Well-foundedness **)
|
|
229 |
|
|
230 |
goal OrderArith.thy
|
|
231 |
"!!r. [| f: A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))";
|
|
232 |
by (rtac wf_onI2 1);
|
|
233 |
by (subgoal_tac "ALL z:A. f`z=f`y --> z: Ba" 1);
|
|
234 |
by (fast_tac ZF_cs 1);
|
|
235 |
by (eres_inst_tac [("a","f`y")] wf_on_induct 1);
|
|
236 |
by (fast_tac (ZF_cs addSIs [apply_type]) 1);
|
|
237 |
by (best_tac (ZF_cs addSIs [apply_type] addSDs [rvimage_iff RS iffD1]) 1);
|
|
238 |
val wf_on_rvimage = result();
|
|
239 |
|
|
240 |
goal OrderArith.thy
|
|
241 |
"!!r. [| f: inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))";
|
|
242 |
by (rtac well_ordI 1);
|
|
243 |
by (rewrite_goals_tac [well_ord_def, tot_ord_def]);
|
|
244 |
by (fast_tac (ZF_cs addSIs [wf_on_rvimage, inj_is_fun]) 1);
|
|
245 |
by (fast_tac (ZF_cs addSIs [linear_rvimage]) 1);
|
|
246 |
val well_ord_rvimage = result();
|