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