author | paulson |
Thu, 11 Nov 1999 10:24:14 +0100 | |
changeset 8004 | 6273f58ea2c1 |
parent 7570 | a9391550eea1 |
child 9173 | 422968aeed49 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/wf.ML |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Tobias Nipkow and Lawrence C Paulson |
4515 | 4 |
Copyright 1998 University of Cambridge |
0 | 5 |
|
4515 | 6 |
Well-founded Recursion |
0 | 7 |
|
8 |
Derived first for transitive relations, and finally for arbitrary WF relations |
|
9 |
via wf_trancl and trans_trancl. |
|
10 |
||
11 |
It is difficult to derive this general case directly, using r^+ instead of |
|
12 |
r. In is_recfun, the two occurrences of the relation must have the same |
|
13 |
form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with |
|
14 |
r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in |
|
15 |
principle, but harder to use, especially to prove wfrec_eclose_eq in |
|
16 |
epsilon.ML. Expanding out the definition of wftrec in wfrec would yield |
|
17 |
a mess. |
|
18 |
*) |
|
19 |
||
20 |
open WF; |
|
21 |
||
22 |
||
23 |
(*** Well-founded relations ***) |
|
24 |
||
435 | 25 |
(** Equivalences between wf and wf_on **) |
26 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
27 |
Goalw [wf_def, wf_on_def] "wf(r) ==> wf[A](r)"; |
4515 | 28 |
by (Clarify_tac 1); (*essential for Blast_tac's efficiency*) |
3016 | 29 |
by (Blast_tac 1); |
760 | 30 |
qed "wf_imp_wf_on"; |
435 | 31 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
32 |
Goalw [wf_def, wf_on_def] "wf[field(r)](r) ==> wf(r)"; |
2469 | 33 |
by (Fast_tac 1); |
760 | 34 |
qed "wf_on_field_imp_wf"; |
435 | 35 |
|
5067 | 36 |
Goal "wf(r) <-> wf[field(r)](r)"; |
4091 | 37 |
by (blast_tac (claset() addIs [wf_imp_wf_on, wf_on_field_imp_wf]) 1); |
760 | 38 |
qed "wf_iff_wf_on_field"; |
0 | 39 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
40 |
Goalw [wf_on_def, wf_def] "[| wf[A](r); B<=A |] ==> wf[B](r)"; |
5265
9d1d4c43c76d
Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents:
5147
diff
changeset
|
41 |
by (Fast_tac 1); |
760 | 42 |
qed "wf_on_subset_A"; |
435 | 43 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
44 |
Goalw [wf_on_def, wf_def] "[| wf[A](r); s<=r |] ==> wf[A](s)"; |
5265
9d1d4c43c76d
Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents:
5147
diff
changeset
|
45 |
by (Fast_tac 1); |
760 | 46 |
qed "wf_on_subset_r"; |
435 | 47 |
|
48 |
(** Introduction rules for wf_on **) |
|
49 |
||
50 |
(*If every non-empty subset of A has an r-minimal element then wf[A](r).*) |
|
5321 | 51 |
val [prem] = Goalw [wf_on_def, wf_def] |
435 | 52 |
"[| !!Z u. [| Z<=A; u:Z; ALL x:Z. EX y:Z. <y,x>:r |] ==> False |] \ |
53 |
\ ==> wf[A](r)"; |
|
0 | 54 |
by (rtac (equals0I RS disjCI RS allI) 1); |
435 | 55 |
by (res_inst_tac [ ("Z", "Z") ] prem 1); |
3016 | 56 |
by (ALLGOALS Blast_tac); |
760 | 57 |
qed "wf_onI"; |
0 | 58 |
|
435 | 59 |
(*If r allows well-founded induction over A then wf[A](r) |
60 |
Premise is equivalent to |
|
61 |
!!B. ALL x:A. (ALL y. <y,x>: r --> y:B) --> x:B ==> A<=B *) |
|
5321 | 62 |
val [prem] = Goal |
435 | 63 |
"[| !!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B; y:A \ |
64 |
\ |] ==> y:B |] \ |
|
65 |
\ ==> wf[A](r)"; |
|
437 | 66 |
by (rtac wf_onI 1); |
435 | 67 |
by (res_inst_tac [ ("c", "u") ] (prem RS DiffE) 1); |
68 |
by (contr_tac 3); |
|
3016 | 69 |
by (Blast_tac 2); |
2469 | 70 |
by (Fast_tac 1); |
760 | 71 |
qed "wf_onI2"; |
0 | 72 |
|
73 |
||
74 |
(** Well-founded Induction **) |
|
75 |
||
76 |
(*Consider the least z in domain(r) Un {a} such that P(z) does not hold...*) |
|
5321 | 77 |
val [major,minor] = Goalw [wf_def] |
0 | 78 |
"[| wf(r); \ |
79 |
\ !!x.[| ALL y. <y,x>: r --> P(y) |] ==> P(x) \ |
|
80 |
\ |] ==> P(a)"; |
|
81 |
by (res_inst_tac [ ("x", "{z:domain(r) Un {a}. ~P(z)}") ] (major RS allE) 1); |
|
82 |
by (etac disjE 1); |
|
4091 | 83 |
by (blast_tac (claset() addEs [equalityE]) 1); |
84 |
by (asm_full_simp_tac (simpset() addsimps [domainI]) 1); |
|
85 |
by (blast_tac (claset() addSDs [minor]) 1); |
|
760 | 86 |
qed "wf_induct"; |
0 | 87 |
|
88 |
(*Perform induction on i, then prove the wf(r) subgoal using prems. *) |
|
89 |
fun wf_ind_tac a prems i = |
|
90 |
EVERY [res_inst_tac [("a",a)] wf_induct i, |
|
1461 | 91 |
rename_last_tac a ["1"] (i+1), |
92 |
ares_tac prems i]; |
|
0 | 93 |
|
485 | 94 |
(*The form of this rule is designed to match wfI*) |
5321 | 95 |
val wfr::amem::prems = Goal |
0 | 96 |
"[| wf(r); a:A; field(r)<=A; \ |
97 |
\ !!x.[| x: A; ALL y. <y,x>: r --> P(y) |] ==> P(x) \ |
|
98 |
\ |] ==> P(a)"; |
|
99 |
by (rtac (amem RS rev_mp) 1); |
|
100 |
by (wf_ind_tac "a" [wfr] 1); |
|
101 |
by (rtac impI 1); |
|
102 |
by (eresolve_tac prems 1); |
|
4091 | 103 |
by (blast_tac (claset() addIs (prems RL [subsetD])) 1); |
760 | 104 |
qed "wf_induct2"; |
0 | 105 |
|
5321 | 106 |
Goal "!!r A. field(r Int A*A) <= A"; |
3016 | 107 |
by (Blast_tac 1); |
760 | 108 |
qed "field_Int_square"; |
435 | 109 |
|
5321 | 110 |
val wfr::amem::prems = Goalw [wf_on_def] |
1461 | 111 |
"[| wf[A](r); a:A; \ |
112 |
\ !!x.[| x: A; ALL y:A. <y,x>: r --> P(y) |] ==> P(x) \ |
|
435 | 113 |
\ |] ==> P(a)"; |
114 |
by (rtac ([wfr, amem, field_Int_square] MRS wf_induct2) 1); |
|
115 |
by (REPEAT (ares_tac prems 1)); |
|
3016 | 116 |
by (Blast_tac 1); |
760 | 117 |
qed "wf_on_induct"; |
435 | 118 |
|
119 |
fun wf_on_ind_tac a prems i = |
|
120 |
EVERY [res_inst_tac [("a",a)] wf_on_induct i, |
|
1461 | 121 |
rename_last_tac a ["1"] (i+2), |
122 |
REPEAT (ares_tac prems i)]; |
|
435 | 123 |
|
124 |
(*If r allows well-founded induction then wf(r)*) |
|
5321 | 125 |
val [subs,indhyp] = Goal |
435 | 126 |
"[| field(r)<=A; \ |
127 |
\ !!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B; y:A \ |
|
128 |
\ |] ==> y:B |] \ |
|
129 |
\ ==> wf(r)"; |
|
437 | 130 |
by (rtac ([wf_onI2, subs] MRS (wf_on_subset_A RS wf_on_field_imp_wf)) 1); |
435 | 131 |
by (REPEAT (ares_tac [indhyp] 1)); |
760 | 132 |
qed "wfI"; |
435 | 133 |
|
134 |
||
135 |
(*** Properties of well-founded relations ***) |
|
136 |
||
5137 | 137 |
Goal "wf(r) ==> <a,a> ~: r"; |
435 | 138 |
by (wf_ind_tac "a" [] 1); |
3016 | 139 |
by (Blast_tac 1); |
760 | 140 |
qed "wf_not_refl"; |
435 | 141 |
|
5452 | 142 |
Goal "wf(r) ==> ALL x. <a,x>:r --> <x,a> ~: r"; |
143 |
by (wf_ind_tac "a" [] 1); |
|
3016 | 144 |
by (Blast_tac 1); |
5452 | 145 |
qed_spec_mp "wf_not_sym"; |
146 |
||
147 |
(* [| wf(r); <a,x> : r; ~P ==> <x,a> : r |] ==> P *) |
|
148 |
bind_thm ("wf_asym", wf_not_sym RS swap); |
|
0 | 149 |
|
5137 | 150 |
Goal "[| wf[A](r); a: A |] ==> <a,a> ~: r"; |
435 | 151 |
by (wf_on_ind_tac "a" [] 1); |
3016 | 152 |
by (Blast_tac 1); |
760 | 153 |
qed "wf_on_not_refl"; |
435 | 154 |
|
5452 | 155 |
Goal "[| wf[A](r); a:A; b:A |] ==> <a,b>:r --> <b,a>~:r"; |
156 |
by (res_inst_tac [("x","b")] bspec 1); |
|
157 |
by (assume_tac 2); |
|
158 |
by (wf_on_ind_tac "a" [] 1); |
|
3016 | 159 |
by (Blast_tac 1); |
5452 | 160 |
qed_spec_mp "wf_on_not_sym"; |
161 |
||
162 |
(* [| wf[A](r); <a,b> : r; a:A; b:A; ~P ==> <b,a> : r |] ==> P *) |
|
163 |
bind_thm ("wf_on_asym", wf_on_not_sym RS swap); |
|
164 |
||
165 |
val prems = |
|
166 |
Goal "[| wf[A](r); <a,b>:r; ~P ==> <b,a>:r; a:A; b:A |] ==> P"; |
|
167 |
by (rtac ccontr 1); |
|
168 |
by (rtac (wf_on_not_sym RS notE) 1); |
|
169 |
by (DEPTH_SOLVE (ares_tac prems 1)); |
|
760 | 170 |
qed "wf_on_asym"; |
435 | 171 |
|
172 |
(*Needed to prove well_ordI. Could also reason that wf[A](r) means |
|
173 |
wf(r Int A*A); thus wf( (r Int A*A)^+ ) and use wf_not_refl *) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
174 |
Goal "[| wf[A](r); <a,b>:r; <b,c>:r; <c,a>:r; a:A; b:A; c:A |] ==> P"; |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
175 |
by (subgoal_tac "ALL y:A. ALL z:A. <a,y>:r --> <y,z>:r --> <z,a>:r --> P" 1); |
435 | 176 |
by (wf_on_ind_tac "a" [] 2); |
3016 | 177 |
by (Blast_tac 2); |
178 |
by (Blast_tac 1); |
|
760 | 179 |
qed "wf_on_chain3"; |
435 | 180 |
|
181 |
||
182 |
(*retains the universal formula for later use!*) |
|
183 |
val bchain_tac = EVERY' [rtac (bspec RS mp), assume_tac, assume_tac ]; |
|
184 |
||
185 |
(*transitive closure of a WF relation is WF provided A is downwards closed*) |
|
186 |
val [wfr,subs] = goal WF.thy |
|
187 |
"[| wf[A](r); r-``A <= A |] ==> wf[A](r^+)"; |
|
437 | 188 |
by (rtac wf_onI2 1); |
435 | 189 |
by (bchain_tac 1); |
190 |
by (eres_inst_tac [("a","y")] (wfr RS wf_on_induct) 1); |
|
191 |
by (cut_facts_tac [subs] 1); |
|
4091 | 192 |
by (blast_tac (claset() addEs [tranclE]) 1); |
760 | 193 |
qed "wf_on_trancl"; |
435 | 194 |
|
5137 | 195 |
Goal "wf(r) ==> wf(r^+)"; |
4091 | 196 |
by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1); |
437 | 197 |
by (rtac (trancl_type RS field_rel_subset RSN (2, wf_on_subset_A)) 1); |
198 |
by (etac wf_on_trancl 1); |
|
3016 | 199 |
by (Blast_tac 1); |
760 | 200 |
qed "wf_trancl"; |
0 | 201 |
|
435 | 202 |
|
203 |
||
0 | 204 |
(** r-``{a} is the set of everything under a in r **) |
205 |
||
6112 | 206 |
bind_thm ("underI", vimage_singleton_iff RS iffD2); |
207 |
bind_thm ("underD", vimage_singleton_iff RS iffD1); |
|
0 | 208 |
|
209 |
(** is_recfun **) |
|
210 |
||
5321 | 211 |
Goalw [is_recfun_def] "is_recfun(r,a,H,f) ==> f: r-``{a} -> range(f)"; |
212 |
by (etac ssubst 1); |
|
0 | 213 |
by (rtac (lamI RS rangeI RS lam_type) 1); |
214 |
by (assume_tac 1); |
|
760 | 215 |
qed "is_recfun_type"; |
0 | 216 |
|
217 |
val [isrec,rel] = goalw WF.thy [is_recfun_def] |
|
218 |
"[| is_recfun(r,a,H,f); <x,a>:r |] ==> f`x = H(x, restrict(f,r-``{x}))"; |
|
443
10884e64c241
added parentheses made necessary by new constrain precedence
clasohm
parents:
437
diff
changeset
|
219 |
by (res_inst_tac [("P", "%x.?t(x) = (?u::i)")] (isrec RS ssubst) 1); |
0 | 220 |
by (rtac (rel RS underI RS beta) 1); |
760 | 221 |
qed "apply_recfun"; |
0 | 222 |
|
223 |
(*eresolve_tac transD solves <a,b>:r using transitivity AT MOST ONCE |
|
224 |
spec RS mp instantiates induction hypotheses*) |
|
225 |
fun indhyp_tac hyps = |
|
6112 | 226 |
resolve_tac (TrueI::refl::reflexive_thm::hyps) ORELSE' |
0 | 227 |
(cut_facts_tac hyps THEN' |
228 |
DEPTH_SOLVE_1 o (ares_tac [TrueI, ballI] ORELSE' |
|
1461 | 229 |
eresolve_tac [underD, transD, spec RS mp])); |
0 | 230 |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
231 |
(*** NOTE! some simplifications need a different solver!! ***) |
7570 | 232 |
val wf_super_ss = simpset() setSolver (mk_solver "WF" indhyp_tac); |
0 | 233 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
234 |
Goalw [is_recfun_def] |
0 | 235 |
"[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] ==> \ |
236 |
\ <x,a>:r --> <x,b>:r --> f`x=g`x"; |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
237 |
by (wf_ind_tac "x" [] 1); |
0 | 238 |
by (REPEAT (rtac impI 1 ORELSE etac ssubst 1)); |
239 |
by (rewtac restrict_def); |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
240 |
by (asm_simp_tac (wf_super_ss addsimps [vimage_singleton_iff]) 1); |
6112 | 241 |
qed_spec_mp "is_recfun_equal"; |
0 | 242 |
|
243 |
val prems as [wfr,transr,recf,recg,_] = goal WF.thy |
|
244 |
"[| wf(r); trans(r); \ |
|
245 |
\ is_recfun(r,a,H,f); is_recfun(r,b,H,g); <b,a>:r |] ==> \ |
|
246 |
\ restrict(f, r-``{b}) = g"; |
|
247 |
by (cut_facts_tac prems 1); |
|
248 |
by (rtac (consI1 RS restrict_type RS fun_extension) 1); |
|
249 |
by (etac is_recfun_type 1); |
|
250 |
by (ALLGOALS |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
251 |
(asm_simp_tac (wf_super_ss addsimps |
1461 | 252 |
[ [wfr,transr,recf,recg] MRS is_recfun_equal ]))); |
760 | 253 |
qed "is_recfun_cut"; |
0 | 254 |
|
255 |
(*** Main Existence Lemma ***) |
|
256 |
||
5321 | 257 |
Goal "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |] ==> f=g"; |
0 | 258 |
by (rtac fun_extension 1); |
259 |
by (REPEAT (ares_tac [is_recfun_equal] 1 |
|
260 |
ORELSE eresolve_tac [is_recfun_type,underD] 1)); |
|
760 | 261 |
qed "is_recfun_functional"; |
0 | 262 |
|
263 |
(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) |
|
5321 | 264 |
Goalw [the_recfun_def] |
0 | 265 |
"[| is_recfun(r,a,H,f); wf(r); trans(r) |] \ |
266 |
\ ==> is_recfun(r, a, H, the_recfun(r,a,H))"; |
|
267 |
by (rtac (ex1I RS theI) 1); |
|
5321 | 268 |
by (REPEAT (ares_tac [is_recfun_functional] 1)); |
760 | 269 |
qed "is_the_recfun"; |
0 | 270 |
|
5321 | 271 |
Goal "[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))"; |
272 |
by (wf_ind_tac "a" [] 1); |
|
0 | 273 |
by (res_inst_tac [("f", "lam y: r-``{a1}. wftrec(r,y,H)")] is_the_recfun 1); |
274 |
by (REPEAT (assume_tac 2)); |
|
275 |
by (rewrite_goals_tac [is_recfun_def, wftrec_def]); |
|
276 |
(*Applying the substitution: must keep the quantified assumption!!*) |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
277 |
by (REPEAT (dtac underD 1 ORELSE resolve_tac [refl, lam_cong] 1)); |
0 | 278 |
by (fold_tac [is_recfun_def]); |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
279 |
by (rtac (consI1 RS restrict_type RSN (2,fun_extension) RS subst_context) 1); |
0 | 280 |
by (rtac is_recfun_type 1); |
281 |
by (ALLGOALS |
|
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
282 |
(asm_simp_tac |
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
283 |
(wf_super_ss addsimps [underI RS beta, apply_recfun, is_recfun_cut]))); |
760 | 284 |
qed "unfold_the_recfun"; |
0 | 285 |
|
286 |
||
287 |
(*** Unfolding wftrec ***) |
|
288 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
289 |
Goal "[| wf(r); trans(r); <b,a>:r |] ==> \ |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
290 |
\ restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)"; |
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
291 |
by (REPEAT (ares_tac [is_recfun_cut, unfold_the_recfun] 1)); |
760 | 292 |
qed "the_recfun_cut"; |
0 | 293 |
|
4515 | 294 |
(*NOT SUITABLE FOR REWRITING: it is recursive!*) |
5067 | 295 |
Goalw [wftrec_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
296 |
"[| wf(r); trans(r) |] ==> \ |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
297 |
\ wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))"; |
2033 | 298 |
by (stac (rewrite_rule [is_recfun_def] unfold_the_recfun) 1); |
4515 | 299 |
by (ALLGOALS |
300 |
(asm_simp_tac |
|
301 |
(simpset() addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut]))); |
|
760 | 302 |
qed "wftrec"; |
0 | 303 |
|
304 |
(** Removal of the premise trans(r) **) |
|
305 |
||
4515 | 306 |
(*NOT SUITABLE FOR REWRITING: it is recursive!*) |
0 | 307 |
val [wfr] = goalw WF.thy [wfrec_def] |
308 |
"wf(r) ==> wfrec(r,a,H) = H(a, lam x:r-``{a}. wfrec(r,x,H))"; |
|
2033 | 309 |
by (stac (wfr RS wf_trancl RS wftrec) 1); |
0 | 310 |
by (rtac trans_trancl 1); |
6
8ce8c4d13d4d
Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents:
0
diff
changeset
|
311 |
by (rtac (vimage_pair_mono RS restrict_lam_eq RS subst_context) 1); |
0 | 312 |
by (etac r_into_trancl 1); |
313 |
by (rtac subset_refl 1); |
|
760 | 314 |
qed "wfrec"; |
0 | 315 |
|
316 |
(*This form avoids giant explosions in proofs. NOTE USE OF == *) |
|
5321 | 317 |
val rew::prems = Goal |
0 | 318 |
"[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==> \ |
319 |
\ h(a) = H(a, lam x: r-``{a}. h(x))"; |
|
320 |
by (rewtac rew); |
|
321 |
by (REPEAT (resolve_tac (prems@[wfrec]) 1)); |
|
760 | 322 |
qed "def_wfrec"; |
0 | 323 |
|
5321 | 324 |
val prems = Goal |
0 | 325 |
"[| wf(r); a:A; field(r)<=A; \ |
326 |
\ !!x u. [| x: A; u: Pi(r-``{x}, B) |] ==> H(x,u) : B(x) \ |
|
327 |
\ |] ==> wfrec(r,a,H) : B(a)"; |
|
328 |
by (res_inst_tac [("a","a")] wf_induct2 1); |
|
2033 | 329 |
by (stac wfrec 4); |
0 | 330 |
by (REPEAT (ares_tac (prems@[lam_type]) 1 |
331 |
ORELSE eresolve_tac [spec RS mp, underD] 1)); |
|
760 | 332 |
qed "wfrec_type"; |
435 | 333 |
|
334 |
||
5067 | 335 |
Goalw [wf_on_def, wfrec_on_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
336 |
"[| wf[A](r); a: A |] ==> \ |
435 | 337 |
\ wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))"; |
437 | 338 |
by (etac (wfrec RS trans) 1); |
4091 | 339 |
by (asm_simp_tac (simpset() addsimps [vimage_Int_square, cons_subset_iff]) 1); |
760 | 340 |
qed "wfrec_on"; |
435 | 341 |