author | wenzelm |
Sun, 23 Jul 2000 12:02:22 +0200 | |
changeset 9410 | 612ee826a409 |
parent 9173 | 422968aeed49 |
child 9907 | 473a6604da94 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Univ |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
435 | 4 |
Copyright 1994 University of Cambridge |
0 | 5 |
|
6 |
The cumulative hierarchy and a small universe for recursive types |
|
7 |
*) |
|
8 |
||
9 |
(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) |
|
5067 | 10 |
Goal "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))"; |
2033 | 11 |
by (stac (Vfrom_def RS def_transrec) 1); |
2469 | 12 |
by (Simp_tac 1); |
760 | 13 |
qed "Vfrom"; |
0 | 14 |
|
15 |
(** Monotonicity **) |
|
16 |
||
5137 | 17 |
Goal "A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"; |
0 | 18 |
by (eps_ind_tac "i" 1); |
19 |
by (rtac (impI RS allI) 1); |
|
2033 | 20 |
by (stac Vfrom 1); |
21 |
by (stac Vfrom 1); |
|
0 | 22 |
by (etac Un_mono 1); |
23 |
by (rtac UN_mono 1); |
|
24 |
by (assume_tac 1); |
|
25 |
by (rtac Pow_mono 1); |
|
26 |
by (etac (bspec RS spec RS mp) 1); |
|
27 |
by (assume_tac 1); |
|
28 |
by (rtac subset_refl 1); |
|
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3074
diff
changeset
|
29 |
qed_spec_mp "Vfrom_mono"; |
0 | 30 |
|
31 |
||
32 |
(** A fundamental equality: Vfrom does not require ordinals! **) |
|
33 |
||
5067 | 34 |
Goal "Vfrom(A,x) <= Vfrom(A,rank(x))"; |
0 | 35 |
by (eps_ind_tac "x" 1); |
2033 | 36 |
by (stac Vfrom 1); |
37 |
by (stac Vfrom 1); |
|
4091 | 38 |
by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1); |
760 | 39 |
qed "Vfrom_rank_subset1"; |
0 | 40 |
|
5067 | 41 |
Goal "Vfrom(A,rank(x)) <= Vfrom(A,x)"; |
0 | 42 |
by (eps_ind_tac "x" 1); |
2033 | 43 |
by (stac Vfrom 1); |
44 |
by (stac Vfrom 1); |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
45 |
by (rtac (subset_refl RS Un_mono) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
46 |
by (rtac UN_least 1); |
27 | 47 |
(*expand rank(x1) = (UN y:x1. succ(rank(y))) in assumptions*) |
48 |
by (etac (rank RS equalityD1 RS subsetD RS UN_E) 1); |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
49 |
by (rtac subset_trans 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
50 |
by (etac UN_upper 2); |
27 | 51 |
by (rtac (subset_refl RS Vfrom_mono RS subset_trans RS Pow_mono) 1); |
52 |
by (etac (ltI RS le_imp_subset) 1); |
|
53 |
by (rtac (Ord_rank RS Ord_succ) 1); |
|
0 | 54 |
by (etac bspec 1); |
55 |
by (assume_tac 1); |
|
760 | 56 |
qed "Vfrom_rank_subset2"; |
0 | 57 |
|
5067 | 58 |
Goal "Vfrom(A,rank(x)) = Vfrom(A,x)"; |
0 | 59 |
by (rtac equalityI 1); |
60 |
by (rtac Vfrom_rank_subset2 1); |
|
61 |
by (rtac Vfrom_rank_subset1 1); |
|
760 | 62 |
qed "Vfrom_rank_eq"; |
0 | 63 |
|
64 |
||
65 |
(*** Basic closure properties ***) |
|
66 |
||
5137 | 67 |
Goal "y:x ==> 0 : Vfrom(A,x)"; |
2033 | 68 |
by (stac Vfrom 1); |
2925 | 69 |
by (Blast_tac 1); |
760 | 70 |
qed "zero_in_Vfrom"; |
0 | 71 |
|
5067 | 72 |
Goal "i <= Vfrom(A,i)"; |
0 | 73 |
by (eps_ind_tac "i" 1); |
2033 | 74 |
by (stac Vfrom 1); |
2925 | 75 |
by (Blast_tac 1); |
760 | 76 |
qed "i_subset_Vfrom"; |
0 | 77 |
|
5067 | 78 |
Goal "A <= Vfrom(A,i)"; |
2033 | 79 |
by (stac Vfrom 1); |
0 | 80 |
by (rtac Un_upper1 1); |
760 | 81 |
qed "A_subset_Vfrom"; |
0 | 82 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
83 |
bind_thm ("A_into_Vfrom", A_subset_Vfrom RS subsetD); |
488 | 84 |
|
5137 | 85 |
Goal "a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))"; |
2033 | 86 |
by (stac Vfrom 1); |
2925 | 87 |
by (Blast_tac 1); |
760 | 88 |
qed "subset_mem_Vfrom"; |
0 | 89 |
|
90 |
(** Finite sets and ordered pairs **) |
|
91 |
||
5137 | 92 |
Goal "a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))"; |
0 | 93 |
by (rtac subset_mem_Vfrom 1); |
4152 | 94 |
by Safe_tac; |
760 | 95 |
qed "singleton_in_Vfrom"; |
0 | 96 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
97 |
Goal "[| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))"; |
0 | 98 |
by (rtac subset_mem_Vfrom 1); |
4152 | 99 |
by Safe_tac; |
760 | 100 |
qed "doubleton_in_Vfrom"; |
0 | 101 |
|
5067 | 102 |
Goalw [Pair_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
103 |
"[| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> \ |
0 | 104 |
\ <a,b> : Vfrom(A,succ(succ(i)))"; |
105 |
by (REPEAT (ares_tac [doubleton_in_Vfrom] 1)); |
|
760 | 106 |
qed "Pair_in_Vfrom"; |
0 | 107 |
|
5321 | 108 |
Goal "a<=Vfrom(A,i) ==> succ(a) : Vfrom(A,succ(succ(i)))"; |
0 | 109 |
by (REPEAT (resolve_tac [subset_mem_Vfrom, succ_subsetI] 1)); |
110 |
by (rtac (Vfrom_mono RSN (2,subset_trans)) 2); |
|
5321 | 111 |
by (REPEAT (ares_tac [subset_refl, subset_succI] 1)); |
760 | 112 |
qed "succ_in_Vfrom"; |
0 | 113 |
|
114 |
(*** 0, successor and limit equations fof Vfrom ***) |
|
115 |
||
5067 | 116 |
Goal "Vfrom(A,0) = A"; |
2033 | 117 |
by (stac Vfrom 1); |
2925 | 118 |
by (Blast_tac 1); |
760 | 119 |
qed "Vfrom_0"; |
0 | 120 |
|
5137 | 121 |
Goal "Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; |
0 | 122 |
by (rtac (Vfrom RS trans) 1); |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
123 |
by (rtac (succI1 RS RepFunI RS Union_upper RSN |
1461 | 124 |
(2, equalityI RS subst_context)) 1); |
0 | 125 |
by (rtac UN_least 1); |
126 |
by (rtac (subset_refl RS Vfrom_mono RS Pow_mono) 1); |
|
27 | 127 |
by (etac (ltI RS le_imp_subset) 1); |
128 |
by (etac Ord_succ 1); |
|
760 | 129 |
qed "Vfrom_succ_lemma"; |
0 | 130 |
|
5067 | 131 |
Goal "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; |
0 | 132 |
by (res_inst_tac [("x1", "succ(i)")] (Vfrom_rank_eq RS subst) 1); |
133 |
by (res_inst_tac [("x1", "i")] (Vfrom_rank_eq RS subst) 1); |
|
2033 | 134 |
by (stac rank_succ 1); |
0 | 135 |
by (rtac (Ord_rank RS Vfrom_succ_lemma) 1); |
760 | 136 |
qed "Vfrom_succ"; |
0 | 137 |
|
138 |
(*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces |
|
139 |
the conclusion to be Vfrom(A,Union(X)) = A Un (UN y:X. Vfrom(A,y)) *) |
|
5321 | 140 |
Goal "y:X ==> Vfrom(A,Union(X)) = (UN y:X. Vfrom(A,y))"; |
141 |
by (stac Vfrom 1); |
|
142 |
by (rtac equalityI 1); |
|
143 |
(*first inclusion*) |
|
144 |
by (rtac Un_least 1); |
|
145 |
by (rtac (A_subset_Vfrom RS subset_trans) 1); |
|
146 |
by (rtac UN_upper 1); |
|
147 |
by (assume_tac 1); |
|
148 |
by (rtac UN_least 1); |
|
149 |
by (etac UnionE 1); |
|
150 |
by (rtac subset_trans 1); |
|
151 |
by (etac UN_upper 2 THEN stac Vfrom 1 THEN |
|
152 |
etac ([UN_upper, Un_upper2] MRS subset_trans) 1); |
|
153 |
(*opposite inclusion*) |
|
154 |
by (rtac UN_least 1); |
|
155 |
by (stac Vfrom 1); |
|
156 |
by (Blast_tac 1); |
|
157 |
qed "Vfrom_Union"; |
|
158 |
||
0 | 159 |
val [prem] = goal Univ.thy "y:X ==> Vfrom(A,Union(X)) = (UN y:X. Vfrom(A,y))"; |
2033 | 160 |
by (stac Vfrom 1); |
0 | 161 |
by (rtac equalityI 1); |
162 |
(*first inclusion*) |
|
163 |
by (rtac Un_least 1); |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
164 |
by (rtac (A_subset_Vfrom RS subset_trans) 1); |
5321 | 165 |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
166 |
by (rtac (prem RS UN_upper) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
167 |
by (rtac UN_least 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
168 |
by (etac UnionE 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
169 |
by (rtac subset_trans 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
170 |
by (etac UN_upper 2); |
2033 | 171 |
by (stac Vfrom 1); |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
172 |
by (etac ([UN_upper, Un_upper2] MRS subset_trans) 1); |
0 | 173 |
(*opposite inclusion*) |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
174 |
by (rtac UN_least 1); |
2033 | 175 |
by (stac Vfrom 1); |
2925 | 176 |
by (Blast_tac 1); |
760 | 177 |
qed "Vfrom_Union"; |
0 | 178 |
|
179 |
(*** Vfrom applied to Limit ordinals ***) |
|
180 |
||
181 |
(*NB. limit ordinals are non-empty; |
|
182 |
Vfrom(A,0) = A = A Un (UN y:0. Vfrom(A,y)) *) |
|
183 |
val [limiti] = goal Univ.thy |
|
184 |
"Limit(i) ==> Vfrom(A,i) = (UN y:i. Vfrom(A,y))"; |
|
27 | 185 |
by (rtac (limiti RS (Limit_has_0 RS ltD) RS Vfrom_Union RS subst) 1); |
2033 | 186 |
by (stac (limiti RS Limit_Union_eq) 1); |
0 | 187 |
by (rtac refl 1); |
760 | 188 |
qed "Limit_Vfrom_eq"; |
0 | 189 |
|
5137 | 190 |
Goal "[| a: Vfrom(A,j); Limit(i); j<i |] ==> a : Vfrom(A,i)"; |
27 | 191 |
by (rtac (Limit_Vfrom_eq RS equalityD2 RS subsetD) 1); |
192 |
by (REPEAT (ares_tac [ltD RS UN_I] 1)); |
|
760 | 193 |
qed "Limit_VfromI"; |
27 | 194 |
|
5321 | 195 |
val prems = Goal |
9173 | 196 |
"[| a: Vfrom(A,i); ~R ==> Limit(i); \ |
1461 | 197 |
\ !!x. [| x<i; a: Vfrom(A,x) |] ==> R \ |
27 | 198 |
\ |] ==> R"; |
9173 | 199 |
by (rtac classical 1); |
27 | 200 |
by (rtac (Limit_Vfrom_eq RS equalityD1 RS subsetD RS UN_E) 1); |
201 |
by (REPEAT (ares_tac (prems @ [ltI, Limit_is_Ord]) 1)); |
|
760 | 202 |
qed "Limit_VfromE"; |
0 | 203 |
|
516 | 204 |
val zero_in_VLimit = Limit_has_0 RS ltD RS zero_in_Vfrom; |
484 | 205 |
|
0 | 206 |
val [major,limiti] = goal Univ.thy |
207 |
"[| a: Vfrom(A,i); Limit(i) |] ==> {a} : Vfrom(A,i)"; |
|
27 | 208 |
by (rtac ([major,limiti] MRS Limit_VfromE) 1); |
209 |
by (etac ([singleton_in_Vfrom, limiti] MRS Limit_VfromI) 1); |
|
0 | 210 |
by (etac (limiti RS Limit_has_succ) 1); |
760 | 211 |
qed "singleton_in_VLimit"; |
0 | 212 |
|
9173 | 213 |
bind_thm ("Vfrom_UnI1", Un_upper1 RS (subset_refl RS Vfrom_mono RS subsetD)); |
214 |
bind_thm ("Vfrom_UnI2", Un_upper2 RS (subset_refl RS Vfrom_mono RS subsetD)); |
|
0 | 215 |
|
216 |
(*Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)*) |
|
217 |
val [aprem,bprem,limiti] = goal Univ.thy |
|
218 |
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \ |
|
219 |
\ {a,b} : Vfrom(A,i)"; |
|
27 | 220 |
by (rtac ([aprem,limiti] MRS Limit_VfromE) 1); |
221 |
by (rtac ([bprem,limiti] MRS Limit_VfromE) 1); |
|
222 |
by (rtac ([doubleton_in_Vfrom, limiti] MRS Limit_VfromI) 1); |
|
223 |
by (etac Vfrom_UnI1 1); |
|
224 |
by (etac Vfrom_UnI2 1); |
|
225 |
by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1)); |
|
760 | 226 |
qed "doubleton_in_VLimit"; |
0 | 227 |
|
228 |
val [aprem,bprem,limiti] = goal Univ.thy |
|
229 |
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \ |
|
230 |
\ <a,b> : Vfrom(A,i)"; |
|
231 |
(*Infer that a, b occur at ordinals x,xa < i.*) |
|
27 | 232 |
by (rtac ([aprem,limiti] MRS Limit_VfromE) 1); |
233 |
by (rtac ([bprem,limiti] MRS Limit_VfromE) 1); |
|
234 |
by (rtac ([Pair_in_Vfrom, limiti] MRS Limit_VfromI) 1); |
|
0 | 235 |
(*Infer that succ(succ(x Un xa)) < i *) |
27 | 236 |
by (etac Vfrom_UnI1 1); |
237 |
by (etac Vfrom_UnI2 1); |
|
238 |
by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1)); |
|
760 | 239 |
qed "Pair_in_VLimit"; |
484 | 240 |
|
5137 | 241 |
Goal "Limit(i) ==> Vfrom(A,i)*Vfrom(A,i) <= Vfrom(A,i)"; |
516 | 242 |
by (REPEAT (ares_tac [subsetI,Pair_in_VLimit] 1 |
484 | 243 |
ORELSE eresolve_tac [SigmaE, ssubst] 1)); |
760 | 244 |
qed "product_VLimit"; |
484 | 245 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
246 |
bind_thm ("Sigma_subset_VLimit", |
1461 | 247 |
[Sigma_mono, product_VLimit] MRS subset_trans); |
484 | 248 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
249 |
bind_thm ("nat_subset_VLimit", |
1461 | 250 |
[nat_le_Limit RS le_imp_subset, i_subset_Vfrom] MRS subset_trans); |
484 | 251 |
|
5137 | 252 |
Goal "[| n: nat; Limit(i) |] ==> n : Vfrom(A,i)"; |
516 | 253 |
by (REPEAT (ares_tac [nat_subset_VLimit RS subsetD] 1)); |
760 | 254 |
qed "nat_into_VLimit"; |
484 | 255 |
|
256 |
(** Closure under disjoint union **) |
|
257 |
||
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
258 |
bind_thm ("zero_in_VLimit", Limit_has_0 RS ltD RS zero_in_Vfrom); |
484 | 259 |
|
5137 | 260 |
Goal "Limit(i) ==> 1 : Vfrom(A,i)"; |
516 | 261 |
by (REPEAT (ares_tac [nat_into_VLimit, nat_0I, nat_succI] 1)); |
760 | 262 |
qed "one_in_VLimit"; |
484 | 263 |
|
5067 | 264 |
Goalw [Inl_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
265 |
"[| a: Vfrom(A,i); Limit(i) |] ==> Inl(a) : Vfrom(A,i)"; |
516 | 266 |
by (REPEAT (ares_tac [zero_in_VLimit, Pair_in_VLimit] 1)); |
760 | 267 |
qed "Inl_in_VLimit"; |
484 | 268 |
|
5067 | 269 |
Goalw [Inr_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
270 |
"[| b: Vfrom(A,i); Limit(i) |] ==> Inr(b) : Vfrom(A,i)"; |
516 | 271 |
by (REPEAT (ares_tac [one_in_VLimit, Pair_in_VLimit] 1)); |
760 | 272 |
qed "Inr_in_VLimit"; |
484 | 273 |
|
5137 | 274 |
Goal "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)"; |
4091 | 275 |
by (blast_tac (claset() addSIs [Inl_in_VLimit, Inr_in_VLimit]) 1); |
760 | 276 |
qed "sum_VLimit"; |
484 | 277 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
278 |
bind_thm ("sum_subset_VLimit", [sum_mono, sum_VLimit] MRS subset_trans); |
484 | 279 |
|
0 | 280 |
|
281 |
||
282 |
(*** Properties assuming Transset(A) ***) |
|
283 |
||
5137 | 284 |
Goal "Transset(A) ==> Transset(Vfrom(A,i))"; |
0 | 285 |
by (eps_ind_tac "i" 1); |
2033 | 286 |
by (stac Vfrom 1); |
4091 | 287 |
by (blast_tac (claset() addSIs [Transset_Union_family, Transset_Un, |
1461 | 288 |
Transset_Pow]) 1); |
760 | 289 |
qed "Transset_Vfrom"; |
0 | 290 |
|
5137 | 291 |
Goal "Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"; |
0 | 292 |
by (rtac (Vfrom_succ RS trans) 1); |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
293 |
by (rtac (Un_upper2 RSN (2,equalityI)) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
294 |
by (rtac (subset_refl RSN (2,Un_least)) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
295 |
by (rtac (A_subset_Vfrom RS subset_trans) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
296 |
by (etac (Transset_Vfrom RS (Transset_iff_Pow RS iffD1)) 1); |
760 | 297 |
qed "Transset_Vfrom_succ"; |
0 | 298 |
|
5321 | 299 |
Goalw [Pair_def,Transset_def] "[| <a,b> <= C; Transset(C) |] ==> a: C & b: C"; |
2925 | 300 |
by (Blast_tac 1); |
760 | 301 |
qed "Transset_Pair_subset"; |
0 | 302 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
303 |
Goal "[| <a,b> <= Vfrom(A,i); Transset(A); Limit(i) |] ==> \ |
0 | 304 |
\ <a,b> : Vfrom(A,i)"; |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
305 |
by (etac (Transset_Pair_subset RS conjE) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
306 |
by (etac Transset_Vfrom 1); |
516 | 307 |
by (REPEAT (ares_tac [Pair_in_VLimit] 1)); |
760 | 308 |
qed "Transset_Pair_subset_VLimit"; |
0 | 309 |
|
310 |
||
311 |
(*** Closure under product/sum applied to elements -- thus Vfrom(A,i) |
|
312 |
is a model of simple type theory provided A is a transitive set |
|
313 |
and i is a limit ordinal |
|
314 |
***) |
|
315 |
||
187 | 316 |
(*General theorem for membership in Vfrom(A,i) when i is a limit ordinal*) |
5321 | 317 |
val [aprem,bprem,limiti,step] = Goal |
1461 | 318 |
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); \ |
187 | 319 |
\ !!x y j. [| j<i; 1:j; x: Vfrom(A,j); y: Vfrom(A,j) \ |
1461 | 320 |
\ |] ==> EX k. h(x,y): Vfrom(A,k) & k<i |] ==> \ |
187 | 321 |
\ h(a,b) : Vfrom(A,i)"; |
322 |
(*Infer that a, b occur at ordinals x,xa < i.*) |
|
323 |
by (rtac ([aprem,limiti] MRS Limit_VfromE) 1); |
|
324 |
by (rtac ([bprem,limiti] MRS Limit_VfromE) 1); |
|
828 | 325 |
by (res_inst_tac [("j1", "x Un xa Un 2")] (step RS exE) 1); |
6071 | 326 |
by (blast_tac (claset() addIs [Limit_VfromI, limiti]) 5); |
187 | 327 |
by (etac (Vfrom_UnI2 RS Vfrom_UnI1) 4); |
328 |
by (etac (Vfrom_UnI1 RS Vfrom_UnI1) 3); |
|
329 |
by (rtac (succI1 RS UnI2) 2); |
|
330 |
by (REPEAT (ares_tac [limiti, Limit_has_0, Limit_has_succ, Un_least_lt] 1)); |
|
760 | 331 |
qed "in_VLimit"; |
0 | 332 |
|
333 |
(** products **) |
|
334 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
335 |
Goal "[| a: Vfrom(A,j); b: Vfrom(A,j); Transset(A) |] ==> \ |
187 | 336 |
\ a*b : Vfrom(A, succ(succ(succ(j))))"; |
0 | 337 |
by (dtac Transset_Vfrom 1); |
338 |
by (rtac subset_mem_Vfrom 1); |
|
339 |
by (rewtac Transset_def); |
|
4091 | 340 |
by (blast_tac (claset() addIs [Pair_in_Vfrom]) 1); |
760 | 341 |
qed "prod_in_Vfrom"; |
0 | 342 |
|
343 |
val [aprem,bprem,limiti,transset] = goal Univ.thy |
|
344 |
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ |
|
345 |
\ a*b : Vfrom(A,i)"; |
|
516 | 346 |
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1); |
187 | 347 |
by (REPEAT (ares_tac [exI, conjI, prod_in_Vfrom, transset, |
1461 | 348 |
limiti RS Limit_has_succ] 1)); |
760 | 349 |
qed "prod_in_VLimit"; |
0 | 350 |
|
351 |
(** Disjoint sums, aka Quine ordered pairs **) |
|
352 |
||
5067 | 353 |
Goalw [sum_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
354 |
"[| a: Vfrom(A,j); b: Vfrom(A,j); Transset(A); 1:j |] ==> \ |
187 | 355 |
\ a+b : Vfrom(A, succ(succ(succ(j))))"; |
0 | 356 |
by (dtac Transset_Vfrom 1); |
357 |
by (rtac subset_mem_Vfrom 1); |
|
358 |
by (rewtac Transset_def); |
|
4091 | 359 |
by (blast_tac (claset() addIs [zero_in_Vfrom, Pair_in_Vfrom, |
1461 | 360 |
i_subset_Vfrom RS subsetD]) 1); |
760 | 361 |
qed "sum_in_Vfrom"; |
0 | 362 |
|
363 |
val [aprem,bprem,limiti,transset] = goal Univ.thy |
|
364 |
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ |
|
365 |
\ a+b : Vfrom(A,i)"; |
|
516 | 366 |
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1); |
187 | 367 |
by (REPEAT (ares_tac [exI, conjI, sum_in_Vfrom, transset, |
1461 | 368 |
limiti RS Limit_has_succ] 1)); |
760 | 369 |
qed "sum_in_VLimit"; |
0 | 370 |
|
371 |
(** function space! **) |
|
372 |
||
5067 | 373 |
Goalw [Pi_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
374 |
"[| a: Vfrom(A,j); b: Vfrom(A,j); Transset(A) |] ==> \ |
187 | 375 |
\ a->b : Vfrom(A, succ(succ(succ(succ(j)))))"; |
0 | 376 |
by (dtac Transset_Vfrom 1); |
377 |
by (rtac subset_mem_Vfrom 1); |
|
378 |
by (rtac (Collect_subset RS subset_trans) 1); |
|
2033 | 379 |
by (stac Vfrom 1); |
0 | 380 |
by (rtac (subset_trans RS subset_trans) 1); |
381 |
by (rtac Un_upper2 3); |
|
382 |
by (rtac (succI1 RS UN_upper) 2); |
|
383 |
by (rtac Pow_mono 1); |
|
384 |
by (rewtac Transset_def); |
|
4091 | 385 |
by (blast_tac (claset() addIs [Pair_in_Vfrom]) 1); |
760 | 386 |
qed "fun_in_Vfrom"; |
0 | 387 |
|
388 |
val [aprem,bprem,limiti,transset] = goal Univ.thy |
|
389 |
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \ |
|
390 |
\ a->b : Vfrom(A,i)"; |
|
516 | 391 |
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1); |
187 | 392 |
by (REPEAT (ares_tac [exI, conjI, fun_in_Vfrom, transset, |
1461 | 393 |
limiti RS Limit_has_succ] 1)); |
760 | 394 |
qed "fun_in_VLimit"; |
0 | 395 |
|
5067 | 396 |
Goalw [Pi_def] |
5321 | 397 |
"[| a: Vfrom(A,j); Transset(A) |] ==> Pow(a) : Vfrom(A, succ(succ(j)))"; |
3074 | 398 |
by (dtac Transset_Vfrom 1); |
399 |
by (rtac subset_mem_Vfrom 1); |
|
400 |
by (rewtac Transset_def); |
|
401 |
by (stac Vfrom 1); |
|
402 |
by (Blast_tac 1); |
|
403 |
qed "Pow_in_Vfrom"; |
|
404 |
||
5268 | 405 |
Goal "[| a: Vfrom(A,i); Limit(i); Transset(A) |] ==> Pow(a) : Vfrom(A,i)"; |
5479 | 406 |
by (blast_tac (claset() addEs [Limit_VfromE] |
9173 | 407 |
addIs [Limit_has_succ, Pow_in_Vfrom, Limit_VfromI]) 1); |
3074 | 408 |
qed "Pow_in_VLimit"; |
409 |
||
0 | 410 |
|
411 |
(*** The set Vset(i) ***) |
|
412 |
||
5067 | 413 |
Goal "Vset(i) = (UN j:i. Pow(Vset(j)))"; |
2033 | 414 |
by (stac Vfrom 1); |
2925 | 415 |
by (Blast_tac 1); |
760 | 416 |
qed "Vset"; |
0 | 417 |
|
418 |
val Vset_succ = Transset_0 RS Transset_Vfrom_succ; |
|
419 |
||
420 |
val Transset_Vset = Transset_0 RS Transset_Vfrom; |
|
421 |
||
422 |
(** Characterisation of the elements of Vset(i) **) |
|
423 |
||
5321 | 424 |
Goal "Ord(i) ==> ALL b. b : Vset(i) --> rank(b) < i"; |
425 |
by (etac trans_induct 1); |
|
2033 | 426 |
by (stac Vset 1); |
4152 | 427 |
by Safe_tac; |
2033 | 428 |
by (stac rank 1); |
27 | 429 |
by (rtac UN_succ_least_lt 1); |
2925 | 430 |
by (Blast_tac 2); |
27 | 431 |
by (REPEAT (ares_tac [ltI] 1)); |
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3074
diff
changeset
|
432 |
qed_spec_mp "VsetD"; |
0 | 433 |
|
5321 | 434 |
Goal "Ord(i) ==> ALL b. rank(b) : i --> b : Vset(i)"; |
435 |
by (etac trans_induct 1); |
|
0 | 436 |
by (rtac allI 1); |
2033 | 437 |
by (stac Vset 1); |
4091 | 438 |
by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1); |
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3074
diff
changeset
|
439 |
val lemma = result(); |
0 | 440 |
|
5137 | 441 |
Goal "rank(x)<i ==> x : Vset(i)"; |
27 | 442 |
by (etac ltE 1); |
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3074
diff
changeset
|
443 |
by (etac (lemma RS spec RS mp) 1); |
27 | 444 |
by (assume_tac 1); |
760 | 445 |
qed "VsetI"; |
0 | 446 |
|
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
447 |
(*Merely a lemma for the result following*) |
5137 | 448 |
Goal "Ord(i) ==> b : Vset(i) <-> rank(b) < i"; |
0 | 449 |
by (rtac iffI 1); |
27 | 450 |
by (REPEAT (eresolve_tac [asm_rl, VsetD, VsetI] 1)); |
760 | 451 |
qed "Vset_Ord_rank_iff"; |
0 | 452 |
|
5067 | 453 |
Goal "b : Vset(a) <-> rank(b) < rank(a)"; |
0 | 454 |
by (rtac (Vfrom_rank_eq RS subst) 1); |
455 |
by (rtac (Ord_rank RS Vset_Ord_rank_iff) 1); |
|
760 | 456 |
qed "Vset_rank_iff"; |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
457 |
Addsimps [Vset_rank_iff]; |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
458 |
|
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
459 |
(*This is rank(rank(a)) = rank(a) *) |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
460 |
Addsimps [Ord_rank RS rank_of_Ord]; |
0 | 461 |
|
5137 | 462 |
Goal "Ord(i) ==> rank(Vset(i)) = i"; |
2033 | 463 |
by (stac rank 1); |
0 | 464 |
by (rtac equalityI 1); |
4152 | 465 |
by Safe_tac; |
828 | 466 |
by (EVERY' [rtac UN_I, |
1461 | 467 |
etac (i_subset_Vfrom RS subsetD), |
468 |
etac (Ord_in_Ord RS rank_of_Ord RS ssubst), |
|
469 |
assume_tac, |
|
470 |
rtac succI1] 3); |
|
27 | 471 |
by (REPEAT (eresolve_tac [asm_rl, VsetD RS ltD, Ord_trans] 1)); |
760 | 472 |
qed "rank_Vset"; |
0 | 473 |
|
474 |
(** Lemmas for reasoning about sets in terms of their elements' ranks **) |
|
475 |
||
5067 | 476 |
Goal "a <= Vset(rank(a))"; |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
477 |
by (rtac subsetI 1); |
27 | 478 |
by (etac (rank_lt RS VsetI) 1); |
760 | 479 |
qed "arg_subset_Vset_rank"; |
0 | 480 |
|
5321 | 481 |
val [iprem] = Goal |
0 | 482 |
"[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b"; |
27 | 483 |
by (rtac ([subset_refl, arg_subset_Vset_rank] MRS |
1461 | 484 |
Int_greatest RS subset_trans) 1); |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
485 |
by (rtac (Ord_rank RS iprem) 1); |
760 | 486 |
qed "Int_Vset_subset"; |
0 | 487 |
|
488 |
(** Set up an environment for simplification **) |
|
489 |
||
5067 | 490 |
Goalw [Inl_def] "rank(a) < rank(Inl(a))"; |
3889
59bab7a52b4c
moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents:
3736
diff
changeset
|
491 |
by (rtac rank_pair2 1); |
59bab7a52b4c
moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents:
3736
diff
changeset
|
492 |
qed "rank_Inl"; |
59bab7a52b4c
moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents:
3736
diff
changeset
|
493 |
|
5067 | 494 |
Goalw [Inr_def] "rank(a) < rank(Inr(a))"; |
3889
59bab7a52b4c
moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents:
3736
diff
changeset
|
495 |
by (rtac rank_pair2 1); |
59bab7a52b4c
moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents:
3736
diff
changeset
|
496 |
qed "rank_Inr"; |
59bab7a52b4c
moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents:
3736
diff
changeset
|
497 |
|
0 | 498 |
val rank_rls = [rank_Inl, rank_Inr, rank_pair1, rank_pair2]; |
27 | 499 |
val rank_trans_rls = rank_rls @ (rank_rls RLN (2, [lt_trans])); |
0 | 500 |
|
4091 | 501 |
val rank_ss = simpset() addsimps [VsetI] addsimps rank_trans_rls; |
0 | 502 |
|
503 |
(** Recursion over Vset levels! **) |
|
504 |
||
505 |
(*NOT SUITABLE FOR REWRITING: recursive!*) |
|
5067 | 506 |
Goalw [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"; |
2033 | 507 |
by (stac transrec 1); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
508 |
by (Simp_tac 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
509 |
by (rtac (refl RS lam_cong RS subst_context) 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
510 |
by (Asm_full_simp_tac 1); |
760 | 511 |
qed "Vrec"; |
0 | 512 |
|
513 |
(*This form avoids giant explosions in proofs. NOTE USE OF == *) |
|
5321 | 514 |
val rew::prems = Goal |
0 | 515 |
"[| !!x. h(x)==Vrec(x,H) |] ==> \ |
516 |
\ h(a) = H(a, lam x: Vset(rank(a)). h(x))"; |
|
517 |
by (rewtac rew); |
|
518 |
by (rtac Vrec 1); |
|
760 | 519 |
qed "def_Vrec"; |
0 | 520 |
|
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
521 |
(*NOT SUITABLE FOR REWRITING: recursive!*) |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
522 |
Goalw [Vrecursor_def] |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
523 |
"Vrecursor(H,a) = H(lam x:Vset(rank(a)). Vrecursor(H,x), a)"; |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
524 |
by (stac transrec 1); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
525 |
by (Simp_tac 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
526 |
by (rtac (refl RS lam_cong RS subst_context) 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
6071
diff
changeset
|
527 |
by (Asm_full_simp_tac 1); |
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
528 |
qed "Vrecursor"; |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
529 |
|
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
530 |
(*This form avoids giant explosions in proofs. NOTE USE OF == *) |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
531 |
Goal "h == Vrecursor(H) ==> h(a) = H(lam x: Vset(rank(a)). h(x), a)"; |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
532 |
by (Asm_simp_tac 1); |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
533 |
by (rtac Vrecursor 1); |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
534 |
qed "def_Vrecursor"; |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5479
diff
changeset
|
535 |
|
0 | 536 |
|
537 |
(*** univ(A) ***) |
|
538 |
||
5137 | 539 |
Goalw [univ_def] "A<=B ==> univ(A) <= univ(B)"; |
0 | 540 |
by (etac Vfrom_mono 1); |
541 |
by (rtac subset_refl 1); |
|
760 | 542 |
qed "univ_mono"; |
0 | 543 |
|
5137 | 544 |
Goalw [univ_def] "Transset(A) ==> Transset(univ(A))"; |
0 | 545 |
by (etac Transset_Vfrom 1); |
760 | 546 |
qed "Transset_univ"; |
0 | 547 |
|
548 |
(** univ(A) as a limit **) |
|
549 |
||
5067 | 550 |
Goalw [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))"; |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
551 |
by (rtac (Limit_nat RS Limit_Vfrom_eq) 1); |
760 | 552 |
qed "univ_eq_UN"; |
0 | 553 |
|
5137 | 554 |
Goal "c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))"; |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
555 |
by (rtac (subset_UN_iff_eq RS iffD1) 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
556 |
by (etac (univ_eq_UN RS subst) 1); |
760 | 557 |
qed "subset_univ_eq_Int"; |
0 | 558 |
|
5321 | 559 |
val [aprem, iprem] = Goal |
1461 | 560 |
"[| a <= univ(X); \ |
561 |
\ !!i. i:nat ==> a Int Vfrom(X,i) <= b \ |
|
0 | 562 |
\ |] ==> a <= b"; |
2033 | 563 |
by (stac (aprem RS subset_univ_eq_Int) 1); |
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
564 |
by (rtac UN_least 1); |
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
565 |
by (etac iprem 1); |
760 | 566 |
qed "univ_Int_Vfrom_subset"; |
0 | 567 |
|
5321 | 568 |
val prems = Goal |
0 | 569 |
"[| a <= univ(X); b <= univ(X); \ |
570 |
\ !!i. i:nat ==> a Int Vfrom(X,i) = b Int Vfrom(X,i) \ |
|
571 |
\ |] ==> a = b"; |
|
15
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
572 |
by (rtac equalityI 1); |
0 | 573 |
by (ALLGOALS |
574 |
(resolve_tac (prems RL [univ_Int_Vfrom_subset]) THEN' |
|
575 |
eresolve_tac (prems RL [equalityD1,equalityD2] RL [subset_trans]) THEN' |
|
576 |
rtac Int_lower1)); |
|
760 | 577 |
qed "univ_Int_Vfrom_eq"; |
0 | 578 |
|
579 |
(** Closure properties **) |
|
580 |
||
5067 | 581 |
Goalw [univ_def] "0 : univ(A)"; |
0 | 582 |
by (rtac (nat_0I RS zero_in_Vfrom) 1); |
760 | 583 |
qed "zero_in_univ"; |
0 | 584 |
|
5067 | 585 |
Goalw [univ_def] "A <= univ(A)"; |
0 | 586 |
by (rtac A_subset_Vfrom 1); |
760 | 587 |
qed "A_subset_univ"; |
0 | 588 |
|
589 |
val A_into_univ = A_subset_univ RS subsetD; |
|
590 |
||
591 |
(** Closure under unordered and ordered pairs **) |
|
592 |
||
5137 | 593 |
Goalw [univ_def] "a: univ(A) ==> {a} : univ(A)"; |
516 | 594 |
by (REPEAT (ares_tac [singleton_in_VLimit, Limit_nat] 1)); |
760 | 595 |
qed "singleton_in_univ"; |
0 | 596 |
|
5067 | 597 |
Goalw [univ_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
598 |
"[| a: univ(A); b: univ(A) |] ==> {a,b} : univ(A)"; |
516 | 599 |
by (REPEAT (ares_tac [doubleton_in_VLimit, Limit_nat] 1)); |
760 | 600 |
qed "doubleton_in_univ"; |
0 | 601 |
|
5067 | 602 |
Goalw [univ_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
603 |
"[| a: univ(A); b: univ(A) |] ==> <a,b> : univ(A)"; |
516 | 604 |
by (REPEAT (ares_tac [Pair_in_VLimit, Limit_nat] 1)); |
760 | 605 |
qed "Pair_in_univ"; |
0 | 606 |
|
5067 | 607 |
Goalw [univ_def] "univ(A)*univ(A) <= univ(A)"; |
516 | 608 |
by (rtac (Limit_nat RS product_VLimit) 1); |
760 | 609 |
qed "product_univ"; |
0 | 610 |
|
611 |
||
612 |
(** The natural numbers **) |
|
613 |
||
5067 | 614 |
Goalw [univ_def] "nat <= univ(A)"; |
0 | 615 |
by (rtac i_subset_Vfrom 1); |
760 | 616 |
qed "nat_subset_univ"; |
0 | 617 |
|
618 |
(* n:nat ==> n:univ(A) *) |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
619 |
bind_thm ("nat_into_univ", (nat_subset_univ RS subsetD)); |
0 | 620 |
|
621 |
(** instances for 1 and 2 **) |
|
622 |
||
5067 | 623 |
Goalw [univ_def] "1 : univ(A)"; |
516 | 624 |
by (rtac (Limit_nat RS one_in_VLimit) 1); |
760 | 625 |
qed "one_in_univ"; |
0 | 626 |
|
627 |
(*unused!*) |
|
5067 | 628 |
Goal "2 : univ(A)"; |
0 | 629 |
by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1)); |
760 | 630 |
qed "two_in_univ"; |
0 | 631 |
|
5067 | 632 |
Goalw [bool_def] "bool <= univ(A)"; |
4091 | 633 |
by (blast_tac (claset() addSIs [zero_in_univ,one_in_univ]) 1); |
760 | 634 |
qed "bool_subset_univ"; |
0 | 635 |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
636 |
bind_thm ("bool_into_univ", (bool_subset_univ RS subsetD)); |
0 | 637 |
|
638 |
||
639 |
(** Closure under disjoint union **) |
|
640 |
||
5137 | 641 |
Goalw [univ_def] "a: univ(A) ==> Inl(a) : univ(A)"; |
516 | 642 |
by (etac (Limit_nat RSN (2,Inl_in_VLimit)) 1); |
760 | 643 |
qed "Inl_in_univ"; |
0 | 644 |
|
5137 | 645 |
Goalw [univ_def] "b: univ(A) ==> Inr(b) : univ(A)"; |
516 | 646 |
by (etac (Limit_nat RSN (2,Inr_in_VLimit)) 1); |
760 | 647 |
qed "Inr_in_univ"; |
0 | 648 |
|
5067 | 649 |
Goalw [univ_def] "univ(C)+univ(C) <= univ(C)"; |
516 | 650 |
by (rtac (Limit_nat RS sum_VLimit) 1); |
760 | 651 |
qed "sum_univ"; |
0 | 652 |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
653 |
bind_thm ("sum_subset_univ", [sum_mono, sum_univ] MRS subset_trans); |
484 | 654 |
|
655 |
||
0 | 656 |
(** Closure under binary union -- use Un_least **) |
657 |
(** Closure under Collect -- use (Collect_subset RS subset_trans) **) |
|
658 |
(** Closure under RepFun -- use RepFun_subset **) |
|
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
659 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
660 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
661 |
(*** Finite Branching Closure Properties ***) |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
662 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
663 |
(** Closure under finite powerset **) |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
664 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
665 |
Goal "[| b: Fin(Vfrom(A,i)); Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j<i"; |
1461 | 666 |
by (etac Fin_induct 1); |
4091 | 667 |
by (blast_tac (claset() addSDs [Limit_has_0]) 1); |
4152 | 668 |
by Safe_tac; |
1461 | 669 |
by (etac Limit_VfromE 1); |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
670 |
by (assume_tac 1); |
4091 | 671 |
by (blast_tac (claset() addSIs [Un_least_lt] addIs [Vfrom_UnI1, Vfrom_UnI2]) 1); |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
672 |
val Fin_Vfrom_lemma = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
673 |
|
5137 | 674 |
Goal "Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
675 |
by (rtac subsetI 1); |
1461 | 676 |
by (dtac Fin_Vfrom_lemma 1); |
4152 | 677 |
by Safe_tac; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
678 |
by (resolve_tac [Vfrom RS ssubst] 1); |
4091 | 679 |
by (blast_tac (claset() addSDs [ltD]) 1); |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
680 |
val Fin_VLimit = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
681 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
682 |
bind_thm ("Fin_subset_VLimit", [Fin_mono, Fin_VLimit] MRS subset_trans); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
683 |
|
5067 | 684 |
Goalw [univ_def] "Fin(univ(A)) <= univ(A)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
685 |
by (rtac (Limit_nat RS Fin_VLimit) 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
686 |
val Fin_univ = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
687 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
688 |
(** Closure under finite powers (functions from a fixed natural number) **) |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
689 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
690 |
Goal "[| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) <= Vfrom(A,i)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
691 |
by (eresolve_tac [nat_fun_subset_Fin RS subset_trans] 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
692 |
by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit, |
1461 | 693 |
nat_subset_VLimit, subset_refl] 1)); |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
694 |
val nat_fun_VLimit = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
695 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
696 |
bind_thm ("nat_fun_subset_VLimit", [Pi_mono, nat_fun_VLimit] MRS subset_trans); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
697 |
|
5137 | 698 |
Goalw [univ_def] "n: nat ==> n -> univ(A) <= univ(A)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
699 |
by (etac (Limit_nat RSN (2,nat_fun_VLimit)) 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
700 |
val nat_fun_univ = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
701 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
702 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
703 |
(** Closure under finite function space **) |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
704 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
705 |
(*General but seldom-used version; normally the domain is fixed*) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
706 |
Goal "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) <= Vfrom(A,i)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
707 |
by (resolve_tac [FiniteFun.dom_subset RS subset_trans] 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
708 |
by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit, subset_refl] 1)); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
709 |
val FiniteFun_VLimit1 = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
710 |
|
5067 | 711 |
Goalw [univ_def] "univ(A) -||> univ(A) <= univ(A)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
712 |
by (rtac (Limit_nat RS FiniteFun_VLimit1) 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
713 |
val FiniteFun_univ1 = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
714 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
715 |
(*Version for a fixed domain*) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
716 |
Goal "[| W <= Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) <= Vfrom(A,i)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
717 |
by (eresolve_tac [subset_refl RSN (2, FiniteFun_mono) RS subset_trans] 1); |
1461 | 718 |
by (etac FiniteFun_VLimit1 1); |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
719 |
val FiniteFun_VLimit = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
720 |
|
5067 | 721 |
Goalw [univ_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
722 |
"W <= univ(A) ==> W -||> univ(A) <= univ(A)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
723 |
by (etac (Limit_nat RSN (2, FiniteFun_VLimit)) 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
724 |
val FiniteFun_univ = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
725 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
726 |
Goal "[| f: W -||> univ(A); W <= univ(A) |] ==> f : univ(A)"; |
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
727 |
by (eresolve_tac [FiniteFun_univ RS subsetD] 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
728 |
by (assume_tac 1); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
729 |
val FiniteFun_in_univ = result(); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
730 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
731 |
(*Remove <= from the rule above*) |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
732 |
val FiniteFun_in_univ' = subsetI RSN (2, FiniteFun_in_univ); |
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
733 |
|
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
782
diff
changeset
|
734 |
|
5321 | 735 |
(**** For QUniv. Properties of Vfrom analogous to the "take-lemma" ****) |
736 |
||
737 |
(*** Intersecting a*b with Vfrom... ***) |
|
738 |
||
739 |
(*This version says a, b exist one level down, in the smaller set Vfrom(X,i)*) |
|
740 |
Goal "[| {a,b} : Vfrom(X,succ(i)); Transset(X) |] \ |
|
741 |
\ ==> a: Vfrom(X,i) & b: Vfrom(X,i)"; |
|
742 |
by (dtac (Transset_Vfrom_succ RS equalityD1 RS subsetD RS PowD) 1); |
|
743 |
by (assume_tac 1); |
|
744 |
by (Fast_tac 1); |
|
745 |
qed "doubleton_in_Vfrom_D"; |
|
746 |
||
747 |
(*This weaker version says a, b exist at the same level*) |
|
748 |
bind_thm ("Vfrom_doubleton_D", Transset_Vfrom RS Transset_doubleton_D); |
|
749 |
||
750 |
(** Using only the weaker theorem would prove <a,b> : Vfrom(X,i) |
|
751 |
implies a, b : Vfrom(X,i), which is useless for induction. |
|
752 |
Using only the stronger theorem would prove <a,b> : Vfrom(X,succ(succ(i))) |
|
753 |
implies a, b : Vfrom(X,i), leaving the succ(i) case untreated. |
|
754 |
The combination gives a reduction by precisely one level, which is |
|
755 |
most convenient for proofs. |
|
756 |
**) |
|
757 |
||
758 |
Goalw [Pair_def] |
|
759 |
"[| <a,b> : Vfrom(X,succ(i)); Transset(X) |] \ |
|
760 |
\ ==> a: Vfrom(X,i) & b: Vfrom(X,i)"; |
|
761 |
by (blast_tac (claset() addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1); |
|
762 |
qed "Pair_in_Vfrom_D"; |
|
763 |
||
764 |
Goal "Transset(X) ==> \ |
|
765 |
\ (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))"; |
|
766 |
by (blast_tac (claset() addSDs [Pair_in_Vfrom_D]) 1); |
|
767 |
qed "product_Int_Vfrom_subset"; |
|
768 |