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