| author | paulson | 
| Mon, 28 Dec 1998 16:46:15 +0100 | |
| changeset 6035 | c041fc54ab4c | 
| parent 5809 | bacf85370ce0 | 
| child 6070 | 032babd0120b | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/epsilon.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 5809 | 6 | Epsilon induction and recursion | 
| 0 | 7 | *) | 
| 8 | ||
| 9 | (*** Basic closure properties ***) | |
| 10 | ||
| 5067 | 11 | Goalw [eclose_def] "A <= eclose(A)"; | 
| 0 | 12 | by (rtac (nat_rec_0 RS equalityD2 RS subset_trans) 1); | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 13 | by (rtac (nat_0I RS UN_upper) 1); | 
| 760 | 14 | qed "arg_subset_eclose"; | 
| 0 | 15 | |
| 16 | val arg_into_eclose = arg_subset_eclose RS subsetD; | |
| 17 | ||
| 5067 | 18 | Goalw [eclose_def,Transset_def] "Transset(eclose(A))"; | 
| 0 | 19 | by (rtac (subsetI RS ballI) 1); | 
| 20 | by (etac UN_E 1); | |
| 21 | by (rtac (nat_succI RS UN_I) 1); | |
| 22 | by (assume_tac 1); | |
| 23 | by (etac (nat_rec_succ RS ssubst) 1); | |
| 24 | by (etac UnionI 1); | |
| 25 | by (assume_tac 1); | |
| 760 | 26 | qed "Transset_eclose"; | 
| 0 | 27 | |
| 28 | (* x : eclose(A) ==> x <= eclose(A) *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 29 | bind_thm ("eclose_subset",
 | 
| 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 30 | rewrite_rule [Transset_def] Transset_eclose RS bspec); | 
| 0 | 31 | |
| 32 | (* [| A : eclose(B); c : A |] ==> c : eclose(B) *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 33 | bind_thm ("ecloseD", eclose_subset RS subsetD);
 | 
| 0 | 34 | |
| 35 | val arg_in_eclose_sing = arg_subset_eclose RS singleton_subsetD; | |
| 36 | val arg_into_eclose_sing = arg_in_eclose_sing RS ecloseD; | |
| 37 | ||
| 38 | (* This is epsilon-induction for eclose(A); see also eclose_induct_down... | |
| 39 | [| a: eclose(A); !!x. [| x: eclose(A); ALL y:x. P(y) |] ==> P(x) | |
| 40 | |] ==> P(a) | |
| 41 | *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 42 | bind_thm ("eclose_induct", Transset_eclose RSN (2, Transset_induct));
 | 
| 0 | 43 | |
| 44 | (*Epsilon induction*) | |
| 45 | val prems = goal Epsilon.thy | |
| 46 | "[| !!x. ALL y:x. P(y) ==> P(x) |] ==> P(a)"; | |
| 47 | by (rtac (arg_in_eclose_sing RS eclose_induct) 1); | |
| 48 | by (eresolve_tac prems 1); | |
| 760 | 49 | qed "eps_induct"; | 
| 0 | 50 | |
| 51 | (*Perform epsilon-induction on i. *) | |
| 52 | fun eps_ind_tac a = | |
| 53 |     EVERY' [res_inst_tac [("a",a)] eps_induct,
 | |
| 1461 | 54 | rename_last_tac a ["1"]]; | 
| 0 | 55 | |
| 56 | ||
| 57 | (*** Leastness of eclose ***) | |
| 58 | ||
| 59 | (** eclose(A) is the least transitive set including A as a subset. **) | |
| 60 | ||
| 5067 | 61 | Goalw [Transset_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 62 | "[| Transset(X); A<=X; n: nat |] ==> \ | 
| 0 | 63 | \ nat_rec(n, A, %m r. Union(r)) <= X"; | 
| 64 | by (etac nat_induct 1); | |
| 4091 | 65 | by (asm_simp_tac (simpset() addsimps [nat_rec_0]) 1); | 
| 66 | by (asm_simp_tac (simpset() addsimps [nat_rec_succ]) 1); | |
| 3016 | 67 | by (Blast_tac 1); | 
| 760 | 68 | qed "eclose_least_lemma"; | 
| 0 | 69 | |
| 5067 | 70 | Goalw [eclose_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 71 | "[| Transset(X); A<=X |] ==> eclose(A) <= X"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 72 | by (rtac (eclose_least_lemma RS UN_least) 1); | 
| 0 | 73 | by (REPEAT (assume_tac 1)); | 
| 760 | 74 | qed "eclose_least"; | 
| 0 | 75 | |
| 76 | (*COMPLETELY DIFFERENT induction principle from eclose_induct!!*) | |
| 77 | val [major,base,step] = goal Epsilon.thy | |
| 1461 | 78 | "[| a: eclose(b); \ | 
| 79 | \ !!y. [| y: b |] ==> P(y); \ | |
| 80 | \ !!y z. [| y: eclose(b); P(y); z: y |] ==> P(z) \ | |
| 0 | 81 | \ |] ==> P(a)"; | 
| 82 | by (rtac (major RSN (3, eclose_least RS subsetD RS CollectD2)) 1); | |
| 83 | by (rtac (CollectI RS subsetI) 2); | |
| 84 | by (etac (arg_subset_eclose RS subsetD) 2); | |
| 85 | by (etac base 2); | |
| 86 | by (rewtac Transset_def); | |
| 4091 | 87 | by (blast_tac (claset() addIs [step,ecloseD]) 1); | 
| 760 | 88 | qed "eclose_induct_down"; | 
| 0 | 89 | |
| 5137 | 90 | Goal "Transset(X) ==> eclose(X) = X"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 91 | by (etac ([eclose_least, arg_subset_eclose] MRS equalityI) 1); | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 92 | by (rtac subset_refl 1); | 
| 760 | 93 | qed "Transset_eclose_eq_arg"; | 
| 0 | 94 | |
| 95 | ||
| 96 | (*** Epsilon recursion ***) | |
| 97 | ||
| 98 | (*Unused...*) | |
| 5137 | 99 | Goal "[| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)"; | 
| 0 | 100 | by (rtac ([Transset_eclose, eclose_subset] MRS eclose_least RS subsetD) 1); | 
| 101 | by (REPEAT (assume_tac 1)); | |
| 760 | 102 | qed "mem_eclose_trans"; | 
| 0 | 103 | |
| 104 | (*Variant of the previous lemma in a useable form for the sequel*) | |
| 5268 | 105 | Goal "[| A: eclose({B});  B: eclose({C}) |] ==> A: eclose({C})";
 | 
| 0 | 106 | by (rtac ([Transset_eclose, singleton_subsetI] MRS eclose_least RS subsetD) 1); | 
| 107 | by (REPEAT (assume_tac 1)); | |
| 760 | 108 | qed "mem_eclose_sing_trans"; | 
| 0 | 109 | |
| 5067 | 110 | Goalw [Transset_def] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 111 |     "[| Transset(i);  j:i |] ==> Memrel(i)-``{j} = j";
 | 
| 4091 | 112 | by (blast_tac (claset() addSIs [MemrelI] addSEs [MemrelE]) 1); | 
| 760 | 113 | qed "under_Memrel"; | 
| 0 | 114 | |
| 115 | (* j : eclose(A) ==> Memrel(eclose(A)) -`` j = j *) | |
| 116 | val under_Memrel_eclose = Transset_eclose RS under_Memrel; | |
| 117 | ||
| 2469 | 118 | val wfrec_ssubst = standard (wf_Memrel RS wfrec RS ssubst); | 
| 0 | 119 | |
| 120 | val [kmemj,jmemi] = goal Epsilon.thy | |
| 121 |     "[| k:eclose({j});  j:eclose({i}) |] ==> \
 | |
| 122 | \    wfrec(Memrel(eclose({i})), k, H) = wfrec(Memrel(eclose({j})), k, H)";
 | |
| 123 | by (rtac (kmemj RS eclose_induct) 1); | |
| 124 | by (rtac wfrec_ssubst 1); | |
| 125 | by (rtac wfrec_ssubst 1); | |
| 4091 | 126 | by (asm_simp_tac (simpset() addsimps [under_Memrel_eclose, | 
| 1461 | 127 | jmemi RSN (2,mem_eclose_sing_trans)]) 1); | 
| 760 | 128 | qed "wfrec_eclose_eq"; | 
| 0 | 129 | |
| 130 | val [prem] = goal Epsilon.thy | |
| 131 |     "k: i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)";
 | |
| 132 | by (rtac (arg_in_eclose_sing RS wfrec_eclose_eq) 1); | |
| 133 | by (rtac (prem RS arg_into_eclose_sing) 1); | |
| 760 | 134 | qed "wfrec_eclose_eq2"; | 
| 0 | 135 | |
| 5067 | 136 | Goalw [transrec_def] | 
| 0 | 137 | "transrec(a,H) = H(a, lam x:a. transrec(x,H))"; | 
| 138 | by (rtac wfrec_ssubst 1); | |
| 4091 | 139 | by (simp_tac (simpset() addsimps [wfrec_eclose_eq2, arg_in_eclose_sing, | 
| 1461 | 140 | under_Memrel_eclose]) 1); | 
| 760 | 141 | qed "transrec"; | 
| 0 | 142 | |
| 143 | (*Avoids explosions in proofs; resolve it with a meta-level definition.*) | |
| 144 | val rew::prems = goal Epsilon.thy | |
| 145 | "[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, lam x:a. f(x))"; | |
| 146 | by (rewtac rew); | |
| 147 | by (REPEAT (resolve_tac (prems@[transrec]) 1)); | |
| 760 | 148 | qed "def_transrec"; | 
| 0 | 149 | |
| 150 | val prems = goal Epsilon.thy | |
| 151 |     "[| !!x u. [| x:eclose({a});  u: Pi(x,B) |] ==> H(x,u) : B(x)   \
 | |
| 152 | \ |] ==> transrec(a,H) : B(a)"; | |
| 153 | by (res_inst_tac [("i", "a")] (arg_in_eclose_sing RS eclose_induct) 1);
 | |
| 2033 | 154 | by (stac transrec 1); | 
| 0 | 155 | by (REPEAT (ares_tac (prems @ [lam_type]) 1 ORELSE etac bspec 1)); | 
| 760 | 156 | qed "transrec_type"; | 
| 0 | 157 | |
| 5137 | 158 | Goal "Ord(i) ==> eclose({i}) <= succ(i)";
 | 
| 0 | 159 | by (etac (Ord_is_Transset RS Transset_succ RS eclose_least) 1); | 
| 160 | by (rtac (succI1 RS singleton_subsetI) 1); | |
| 760 | 161 | qed "eclose_sing_Ord"; | 
| 0 | 162 | |
| 163 | val prems = goal Epsilon.thy | |
| 164 | "[| j: i; Ord(i); \ | |
| 165 | \ !!x u. [| x: i; u: Pi(x,B) |] ==> H(x,u) : B(x) \ | |
| 166 | \ |] ==> transrec(j,H) : B(j)"; | |
| 167 | by (rtac transrec_type 1); | |
| 168 | by (resolve_tac prems 1); | |
| 169 | by (rtac (Ord_in_Ord RS eclose_sing_Ord RS subsetD RS succE) 1); | |
| 170 | by (DEPTH_SOLVE (ares_tac prems 1 ORELSE eresolve_tac [ssubst,Ord_trans] 1)); | |
| 760 | 171 | qed "Ord_transrec_type"; | 
| 0 | 172 | |
| 173 | (*** Rank ***) | |
| 174 | ||
| 175 | (*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) | |
| 5067 | 176 | Goal "rank(a) = (UN y:a. succ(rank(y)))"; | 
| 2033 | 177 | by (stac (rank_def RS def_transrec) 1); | 
| 2469 | 178 | by (Simp_tac 1); | 
| 760 | 179 | qed "rank"; | 
| 0 | 180 | |
| 5067 | 181 | Goal "Ord(rank(a))"; | 
| 0 | 182 | by (eps_ind_tac "a" 1); | 
| 2033 | 183 | by (stac rank 1); | 
| 0 | 184 | by (rtac (Ord_succ RS Ord_UN) 1); | 
| 185 | by (etac bspec 1); | |
| 186 | by (assume_tac 1); | |
| 760 | 187 | qed "Ord_rank"; | 
| 0 | 188 | |
| 189 | val [major] = goal Epsilon.thy "Ord(i) ==> rank(i) = i"; | |
| 190 | by (rtac (major RS trans_induct) 1); | |
| 2033 | 191 | by (stac rank 1); | 
| 4091 | 192 | by (asm_simp_tac (simpset() addsimps [Ord_equality]) 1); | 
| 760 | 193 | qed "rank_of_Ord"; | 
| 0 | 194 | |
| 5137 | 195 | Goal "a:b ==> rank(a) < rank(b)"; | 
| 0 | 196 | by (res_inst_tac [("a1","b")] (rank RS ssubst) 1);
 | 
| 129 | 197 | by (etac (UN_I RS ltI) 1); | 
| 0 | 198 | by (rtac succI1 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 199 | by (REPEAT (ares_tac [Ord_UN, Ord_succ, Ord_rank] 1)); | 
| 760 | 200 | qed "rank_lt"; | 
| 0 | 201 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 202 | val [major] = goal Epsilon.thy "a: eclose(b) ==> rank(a) < rank(b)"; | 
| 0 | 203 | by (rtac (major RS eclose_induct_down) 1); | 
| 204 | by (etac rank_lt 1); | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 205 | by (etac (rank_lt RS lt_trans) 1); | 
| 0 | 206 | by (assume_tac 1); | 
| 760 | 207 | qed "eclose_rank_lt"; | 
| 0 | 208 | |
| 5137 | 209 | Goal "a<=b ==> rank(a) le rank(b)"; | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 210 | by (rtac subset_imp_le 1); | 
| 2033 | 211 | by (stac rank 1); | 
| 212 | by (stac rank 1); | |
| 0 | 213 | by (etac UN_mono 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 214 | by (REPEAT (resolve_tac [subset_refl, Ord_rank] 1)); | 
| 760 | 215 | qed "rank_mono"; | 
| 0 | 216 | |
| 5067 | 217 | Goal "rank(Pow(a)) = succ(rank(a))"; | 
| 0 | 218 | by (rtac (rank RS trans) 1); | 
| 437 | 219 | by (rtac le_anti_sym 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 220 | by (DO_GOAL [rtac (Ord_rank RS Ord_succ RS UN_least_le), | 
| 1461 | 221 | etac (PowD RS rank_mono RS succ_leI)] 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 222 | by (DO_GOAL [rtac ([Pow_top, le_refl] MRS UN_upper_le), | 
| 1461 | 223 | REPEAT o rtac (Ord_rank RS Ord_succ)] 1); | 
| 760 | 224 | qed "rank_Pow"; | 
| 0 | 225 | |
| 5067 | 226 | Goal "rank(0) = 0"; | 
| 0 | 227 | by (rtac (rank RS trans) 1); | 
| 3016 | 228 | by (Blast_tac 1); | 
| 760 | 229 | qed "rank_0"; | 
| 0 | 230 | |
| 5067 | 231 | Goal "rank(succ(x)) = succ(rank(x))"; | 
| 0 | 232 | by (rtac (rank RS trans) 1); | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 233 | by (rtac ([UN_least, succI1 RS UN_upper] MRS equalityI) 1); | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 234 | by (etac succE 1); | 
| 3016 | 235 | by (Blast_tac 1); | 
| 129 | 236 | by (etac (rank_lt RS leI RS succ_leI RS le_imp_subset) 1); | 
| 760 | 237 | qed "rank_succ"; | 
| 0 | 238 | |
| 5067 | 239 | Goal "rank(Union(A)) = (UN x:A. rank(x))"; | 
| 0 | 240 | by (rtac equalityI 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 241 | by (rtac (rank_mono RS le_imp_subset RS UN_least) 2); | 
| 0 | 242 | by (etac Union_upper 2); | 
| 2033 | 243 | by (stac rank 1); | 
| 0 | 244 | by (rtac UN_least 1); | 
| 245 | by (etac UnionE 1); | |
| 246 | by (rtac subset_trans 1); | |
| 247 | by (etac (RepFunI RS Union_upper) 2); | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 248 | by (etac (rank_lt RS succ_leI RS le_imp_subset) 1); | 
| 760 | 249 | qed "rank_Union"; | 
| 0 | 250 | |
| 5067 | 251 | Goal "rank(eclose(a)) = rank(a)"; | 
| 437 | 252 | by (rtac le_anti_sym 1); | 
| 0 | 253 | by (rtac (arg_subset_eclose RS rank_mono) 2); | 
| 254 | by (res_inst_tac [("a1","eclose(a)")] (rank RS ssubst) 1);
 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 255 | by (rtac (Ord_rank RS UN_least_le) 1); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 256 | by (etac (eclose_rank_lt RS succ_leI) 1); | 
| 760 | 257 | qed "rank_eclose"; | 
| 0 | 258 | |
| 5067 | 259 | Goalw [Pair_def] "rank(a) < rank(<a,b>)"; | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 260 | by (rtac (consI1 RS rank_lt RS lt_trans) 1); | 
| 0 | 261 | by (rtac (consI1 RS consI2 RS rank_lt) 1); | 
| 760 | 262 | qed "rank_pair1"; | 
| 0 | 263 | |
| 5067 | 264 | Goalw [Pair_def] "rank(b) < rank(<a,b>)"; | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 265 | by (rtac (consI1 RS consI2 RS rank_lt RS lt_trans) 1); | 
| 0 | 266 | by (rtac (consI1 RS consI2 RS rank_lt) 1); | 
| 760 | 267 | qed "rank_pair2"; | 
| 0 | 268 | |
| 269 | (*** Corollaries of leastness ***) | |
| 270 | ||
| 5137 | 271 | Goal "A:B ==> eclose(A)<=eclose(B)"; | 
| 0 | 272 | by (rtac (Transset_eclose RS eclose_least) 1); | 
| 273 | by (etac (arg_into_eclose RS eclose_subset) 1); | |
| 760 | 274 | qed "mem_eclose_subset"; | 
| 0 | 275 | |
| 5137 | 276 | Goal "A<=B ==> eclose(A) <= eclose(B)"; | 
| 0 | 277 | by (rtac (Transset_eclose RS eclose_least) 1); | 
| 278 | by (etac subset_trans 1); | |
| 279 | by (rtac arg_subset_eclose 1); | |
| 760 | 280 | qed "eclose_mono"; | 
| 0 | 281 | |
| 282 | (** Idempotence of eclose **) | |
| 283 | ||
| 5067 | 284 | Goal "eclose(eclose(A)) = eclose(A)"; | 
| 0 | 285 | by (rtac equalityI 1); | 
| 286 | by (rtac ([Transset_eclose, subset_refl] MRS eclose_least) 1); | |
| 287 | by (rtac arg_subset_eclose 1); | |
| 760 | 288 | qed "eclose_idem"; | 
| 2469 | 289 | |
| 5505 | 290 | (*Transfinite recursion for definitions based on the three cases of ordinals*) | 
| 2469 | 291 | |
| 5067 | 292 | Goal "transrec2(0,a,b) = a"; | 
| 2469 | 293 | by (rtac (transrec2_def RS def_transrec RS trans) 1); | 
| 294 | by (Simp_tac 1); | |
| 295 | qed "transrec2_0"; | |
| 296 | ||
| 5067 | 297 | Goal "(THE j. i=j) = i"; | 
| 5505 | 298 | by (Blast_tac 1); | 
| 2469 | 299 | qed "THE_eq"; | 
| 300 | ||
| 5067 | 301 | Goal "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))"; | 
| 2469 | 302 | by (rtac (transrec2_def RS def_transrec RS trans) 1); | 
| 5137 | 303 | by (simp_tac (simpset() addsimps [succ_not_0, THE_eq, if_P]) 1); | 
| 3016 | 304 | by (Blast_tac 1); | 
| 2469 | 305 | qed "transrec2_succ"; | 
| 306 | ||
| 5137 | 307 | Goal "Limit(i) ==> transrec2(i,a,b) = (UN j<i. transrec2(j,a,b))"; | 
| 2469 | 308 | by (rtac (transrec2_def RS def_transrec RS trans) 1); | 
| 5137 | 309 | by (Simp_tac 1); | 
| 4091 | 310 | by (blast_tac (claset() addSDs [Limit_has_0] addSEs [succ_LimitE]) 1); | 
| 2469 | 311 | qed "transrec2_Limit"; | 
| 312 | ||
| 313 | Addsimps [transrec2_0, transrec2_succ]; | |
| 3016 | 314 |