author | obua |
Thu, 16 Feb 2006 04:17:19 +0100 | |
changeset 19067 | c0321d7d6b3d |
parent 16417 | 9bc16273c2d4 |
child 22710 | f44439cdce77 |
permissions | -rw-r--r-- |
13165 | 1 |
(* Title: ZF/WF.thy |
0 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Tobias Nipkow and Lawrence C Paulson |
435 | 4 |
Copyright 1994 University of Cambridge |
0 | 5 |
|
13165 | 6 |
Derived first for transitive relations, and finally for arbitrary WF relations |
7 |
via wf_trancl and trans_trancl. |
|
8 |
||
9 |
It is difficult to derive this general case directly, using r^+ instead of |
|
10 |
r. In is_recfun, the two occurrences of the relation must have the same |
|
11 |
form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with |
|
12 |
r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in |
|
13 |
principle, but harder to use, especially to prove wfrec_eclose_eq in |
|
14 |
epsilon.ML. Expanding out the definition of wftrec in wfrec would yield |
|
15 |
a mess. |
|
0 | 16 |
*) |
17 |
||
13356 | 18 |
header{*Well-Founded Recursion*} |
19 |
||
16417 | 20 |
theory WF imports Trancl begin |
13165 | 21 |
|
22 |
constdefs |
|
23 |
wf :: "i=>o" |
|
24 |
(*r is a well-founded relation*) |
|
25 |
"wf(r) == ALL Z. Z=0 | (EX x:Z. ALL y. <y,x>:r --> ~ y:Z)" |
|
26 |
||
27 |
wf_on :: "[i,i]=>o" ("wf[_]'(_')") |
|
28 |
(*r is well-founded on A*) |
|
29 |
"wf_on(A,r) == wf(r Int A*A)" |
|
30 |
||
31 |
is_recfun :: "[i, i, [i,i]=>i, i] =>o" |
|
32 |
"is_recfun(r,a,H,f) == (f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))))" |
|
33 |
||
34 |
the_recfun :: "[i, i, [i,i]=>i] =>i" |
|
35 |
"the_recfun(r,a,H) == (THE f. is_recfun(r,a,H,f))" |
|
36 |
||
37 |
wftrec :: "[i, i, [i,i]=>i] =>i" |
|
38 |
"wftrec(r,a,H) == H(a, the_recfun(r,a,H))" |
|
39 |
||
40 |
wfrec :: "[i, i, [i,i]=>i] =>i" |
|
41 |
(*public version. Does not require r to be transitive*) |
|
42 |
"wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))" |
|
43 |
||
44 |
wfrec_on :: "[i, i, i, [i,i]=>i] =>i" ("wfrec[_]'(_,_,_')") |
|
45 |
"wfrec[A](r,a,H) == wfrec(r Int A*A, a, H)" |
|
46 |
||
47 |
||
13356 | 48 |
subsection{*Well-Founded Relations*} |
13165 | 49 |
|
13634 | 50 |
subsubsection{*Equivalences between @{term wf} and @{term wf_on}*} |
13165 | 51 |
|
52 |
lemma wf_imp_wf_on: "wf(r) ==> wf[A](r)" |
|
13780 | 53 |
by (unfold wf_def wf_on_def, force) |
13165 | 54 |
|
13248 | 55 |
lemma wf_on_imp_wf: "[|wf[A](r); r <= A*A|] ==> wf(r)"; |
56 |
by (simp add: wf_on_def subset_Int_iff) |
|
57 |
||
13165 | 58 |
lemma wf_on_field_imp_wf: "wf[field(r)](r) ==> wf(r)" |
59 |
by (unfold wf_def wf_on_def, fast) |
|
60 |
||
61 |
lemma wf_iff_wf_on_field: "wf(r) <-> wf[field(r)](r)" |
|
62 |
by (blast intro: wf_imp_wf_on wf_on_field_imp_wf) |
|
63 |
||
64 |
lemma wf_on_subset_A: "[| wf[A](r); B<=A |] ==> wf[B](r)" |
|
65 |
by (unfold wf_on_def wf_def, fast) |
|
66 |
||
67 |
lemma wf_on_subset_r: "[| wf[A](r); s<=r |] ==> wf[A](s)" |
|
68 |
by (unfold wf_on_def wf_def, fast) |
|
69 |
||
13217 | 70 |
lemma wf_subset: "[|wf(s); r<=s|] ==> wf(r)" |
71 |
by (simp add: wf_def, fast) |
|
72 |
||
13634 | 73 |
subsubsection{*Introduction Rules for @{term wf_on}*} |
13165 | 74 |
|
13634 | 75 |
text{*If every non-empty subset of @{term A} has an @{term r}-minimal element |
76 |
then we have @{term "wf[A](r)"}.*} |
|
13165 | 77 |
lemma wf_onI: |
78 |
assumes prem: "!!Z u. [| Z<=A; u:Z; ALL x:Z. EX y:Z. <y,x>:r |] ==> False" |
|
79 |
shows "wf[A](r)" |
|
80 |
apply (unfold wf_on_def wf_def) |
|
81 |
apply (rule equals0I [THEN disjCI, THEN allI]) |
|
13784 | 82 |
apply (rule_tac Z = Z in prem, blast+) |
13165 | 83 |
done |
84 |
||
13634 | 85 |
text{*If @{term r} allows well-founded induction over @{term A} |
86 |
then we have @{term "wf[A](r)"}. Premise is equivalent to |
|
87 |
@{term "!!B. ALL x:A. (ALL y. <y,x>: r --> y:B) --> x:B ==> A<=B"} *} |
|
13165 | 88 |
lemma wf_onI2: |
89 |
assumes prem: "!!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B; y:A |] |
|
90 |
==> y:B" |
|
91 |
shows "wf[A](r)" |
|
92 |
apply (rule wf_onI) |
|
93 |
apply (rule_tac c=u in prem [THEN DiffE]) |
|
94 |
prefer 3 apply blast |
|
95 |
apply fast+ |
|
96 |
done |
|
97 |
||
98 |
||
13634 | 99 |
subsubsection{*Well-founded Induction*} |
13165 | 100 |
|
13634 | 101 |
text{*Consider the least @{term z} in @{term "domain(r)"} such that |
102 |
@{term "P(z)"} does not hold...*} |
|
13534 | 103 |
lemma wf_induct [induct set: wf]: |
13165 | 104 |
"[| wf(r); |
13634 | 105 |
!!x.[| ALL y. <y,x>: r --> P(y) |] ==> P(x) |] |
106 |
==> P(a)" |
|
13165 | 107 |
apply (unfold wf_def) |
13252 | 108 |
apply (erule_tac x = "{z : domain(r). ~ P(z)}" in allE) |
13165 | 109 |
apply blast |
110 |
done |
|
435 | 111 |
|
13534 | 112 |
lemmas wf_induct_rule = wf_induct [rule_format, induct set: wf] |
13203
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13175
diff
changeset
|
113 |
|
13634 | 114 |
text{*The form of this rule is designed to match @{text wfI}*} |
13165 | 115 |
lemma wf_induct2: |
116 |
"[| wf(r); a:A; field(r)<=A; |
|
117 |
!!x.[| x: A; ALL y. <y,x>: r --> P(y) |] ==> P(x) |] |
|
118 |
==> P(a)" |
|
119 |
apply (erule_tac P="a:A" in rev_mp) |
|
120 |
apply (erule_tac a=a in wf_induct, blast) |
|
121 |
done |
|
122 |
||
123 |
lemma field_Int_square: "field(r Int A*A) <= A" |
|
124 |
by blast |
|
125 |
||
13534 | 126 |
lemma wf_on_induct [consumes 2, induct set: wf_on]: |
13165 | 127 |
"[| wf[A](r); a:A; |
128 |
!!x.[| x: A; ALL y:A. <y,x>: r --> P(y) |] ==> P(x) |
|
129 |
|] ==> P(a)" |
|
130 |
apply (unfold wf_on_def) |
|
131 |
apply (erule wf_induct2, assumption) |
|
132 |
apply (rule field_Int_square, blast) |
|
133 |
done |
|
134 |
||
13534 | 135 |
lemmas wf_on_induct_rule = |
136 |
wf_on_induct [rule_format, consumes 2, induct set: wf_on] |
|
13203
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13175
diff
changeset
|
137 |
|
fac77a839aa2
Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents:
13175
diff
changeset
|
138 |
|
13634 | 139 |
text{*If @{term r} allows well-founded induction |
140 |
then we have @{term "wf(r)"}.*} |
|
13165 | 141 |
lemma wfI: |
142 |
"[| field(r)<=A; |
|
143 |
!!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B; y:A|] |
|
144 |
==> y:B |] |
|
145 |
==> wf(r)" |
|
146 |
apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf]) |
|
147 |
apply (rule wf_onI2) |
|
148 |
prefer 2 apply blast |
|
149 |
apply blast |
|
150 |
done |
|
151 |
||
152 |
||
13356 | 153 |
subsection{*Basic Properties of Well-Founded Relations*} |
13165 | 154 |
|
155 |
lemma wf_not_refl: "wf(r) ==> <a,a> ~: r" |
|
156 |
by (erule_tac a=a in wf_induct, blast) |
|
157 |
||
158 |
lemma wf_not_sym [rule_format]: "wf(r) ==> ALL x. <a,x>:r --> <x,a> ~: r" |
|
159 |
by (erule_tac a=a in wf_induct, blast) |
|
160 |
||
161 |
(* [| wf(r); <a,x> : r; ~P ==> <x,a> : r |] ==> P *) |
|
162 |
lemmas wf_asym = wf_not_sym [THEN swap, standard] |
|
163 |
||
164 |
lemma wf_on_not_refl: "[| wf[A](r); a: A |] ==> <a,a> ~: r" |
|
13269 | 165 |
by (erule_tac a=a in wf_on_induct, assumption, blast) |
0 | 166 |
|
13165 | 167 |
lemma wf_on_not_sym [rule_format]: |
168 |
"[| wf[A](r); a:A |] ==> ALL b:A. <a,b>:r --> <b,a>~:r" |
|
13269 | 169 |
apply (erule_tac a=a in wf_on_induct, assumption, blast) |
13165 | 170 |
done |
171 |
||
172 |
lemma wf_on_asym: |
|
173 |
"[| wf[A](r); ~Z ==> <a,b> : r; |
|
174 |
<b,a> ~: r ==> Z; ~Z ==> a : A; ~Z ==> b : A |] ==> Z" |
|
13269 | 175 |
by (blast dest: wf_on_not_sym) |
13165 | 176 |
|
177 |
||
178 |
(*Needed to prove well_ordI. Could also reason that wf[A](r) means |
|
179 |
wf(r Int A*A); thus wf( (r Int A*A)^+ ) and use wf_not_refl *) |
|
180 |
lemma wf_on_chain3: |
|
181 |
"[| wf[A](r); <a,b>:r; <b,c>:r; <c,a>:r; a:A; b:A; c:A |] ==> P" |
|
182 |
apply (subgoal_tac "ALL y:A. ALL z:A. <a,y>:r --> <y,z>:r --> <z,a>:r --> P", |
|
183 |
blast) |
|
13269 | 184 |
apply (erule_tac a=a in wf_on_induct, assumption, blast) |
13165 | 185 |
done |
186 |
||
187 |
||
188 |
||
13634 | 189 |
text{*transitive closure of a WF relation is WF provided |
190 |
@{term A} is downward closed*} |
|
13165 | 191 |
lemma wf_on_trancl: |
192 |
"[| wf[A](r); r-``A <= A |] ==> wf[A](r^+)" |
|
193 |
apply (rule wf_onI2) |
|
194 |
apply (frule bspec [THEN mp], assumption+) |
|
13784 | 195 |
apply (erule_tac a = y in wf_on_induct, assumption) |
13165 | 196 |
apply (blast elim: tranclE, blast) |
197 |
done |
|
198 |
||
199 |
lemma wf_trancl: "wf(r) ==> wf(r^+)" |
|
200 |
apply (simp add: wf_iff_wf_on_field) |
|
201 |
apply (rule wf_on_subset_A) |
|
202 |
apply (erule wf_on_trancl) |
|
203 |
apply blast |
|
204 |
apply (rule trancl_type [THEN field_rel_subset]) |
|
205 |
done |
|
206 |
||
207 |
||
13634 | 208 |
text{*@{term "r-``{a}"} is the set of everything under @{term a} in @{term r}*} |
13165 | 209 |
|
210 |
lemmas underI = vimage_singleton_iff [THEN iffD2, standard] |
|
211 |
lemmas underD = vimage_singleton_iff [THEN iffD1, standard] |
|
212 |
||
13634 | 213 |
|
214 |
subsection{*The Predicate @{term is_recfun}*} |
|
0 | 215 |
|
13165 | 216 |
lemma is_recfun_type: "is_recfun(r,a,H,f) ==> f: r-``{a} -> range(f)" |
217 |
apply (unfold is_recfun_def) |
|
218 |
apply (erule ssubst) |
|
219 |
apply (rule lamI [THEN rangeI, THEN lam_type], assumption) |
|
220 |
done |
|
221 |
||
13269 | 222 |
lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function] |
223 |
||
13165 | 224 |
lemma apply_recfun: |
225 |
"[| is_recfun(r,a,H,f); <x,a>:r |] ==> f`x = H(x, restrict(f,r-``{x}))" |
|
226 |
apply (unfold is_recfun_def) |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13165
diff
changeset
|
227 |
txt{*replace f only on the left-hand side*} |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13165
diff
changeset
|
228 |
apply (erule_tac P = "%x.?t(x) = ?u" in ssubst) |
13269 | 229 |
apply (simp add: underI) |
13165 | 230 |
done |
231 |
||
232 |
lemma is_recfun_equal [rule_format]: |
|
233 |
"[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] |
|
234 |
==> <x,a>:r --> <x,b>:r --> f`x=g`x" |
|
13784 | 235 |
apply (frule_tac f = f in is_recfun_type) |
236 |
apply (frule_tac f = g in is_recfun_type) |
|
13165 | 237 |
apply (simp add: is_recfun_def) |
238 |
apply (erule_tac a=x in wf_induct) |
|
239 |
apply (intro impI) |
|
240 |
apply (elim ssubst) |
|
241 |
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) |
|
242 |
apply (rule_tac t = "%z. H (?x,z) " in subst_context) |
|
243 |
apply (subgoal_tac "ALL y : r-``{x}. ALL z. <y,z>:f <-> <y,z>:g") |
|
244 |
apply (blast dest: transD) |
|
245 |
apply (simp add: apply_iff) |
|
246 |
apply (blast dest: transD intro: sym) |
|
247 |
done |
|
248 |
||
249 |
lemma is_recfun_cut: |
|
250 |
"[| wf(r); trans(r); |
|
251 |
is_recfun(r,a,H,f); is_recfun(r,b,H,g); <b,a>:r |] |
|
252 |
==> restrict(f, r-``{b}) = g" |
|
13784 | 253 |
apply (frule_tac f = f in is_recfun_type) |
13165 | 254 |
apply (rule fun_extension) |
255 |
apply (blast dest: transD intro: restrict_type2) |
|
256 |
apply (erule is_recfun_type, simp) |
|
257 |
apply (blast dest: transD intro: is_recfun_equal) |
|
258 |
done |
|
259 |
||
13356 | 260 |
subsection{*Recursion: Main Existence Lemma*} |
435 | 261 |
|
13165 | 262 |
lemma is_recfun_functional: |
263 |
"[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |] ==> f=g" |
|
264 |
by (blast intro: fun_extension is_recfun_type is_recfun_equal) |
|
265 |
||
13248 | 266 |
lemma the_recfun_eq: |
267 |
"[| is_recfun(r,a,H,f); wf(r); trans(r) |] ==> the_recfun(r,a,H) = f" |
|
268 |
apply (unfold the_recfun_def) |
|
269 |
apply (blast intro: is_recfun_functional) |
|
270 |
done |
|
271 |
||
13165 | 272 |
(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) |
273 |
lemma is_the_recfun: |
|
274 |
"[| is_recfun(r,a,H,f); wf(r); trans(r) |] |
|
275 |
==> is_recfun(r, a, H, the_recfun(r,a,H))" |
|
13248 | 276 |
by (simp add: the_recfun_eq) |
13165 | 277 |
|
278 |
lemma unfold_the_recfun: |
|
279 |
"[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))" |
|
280 |
apply (rule_tac a=a in wf_induct, assumption) |
|
281 |
apply (rename_tac a1) |
|
282 |
apply (rule_tac f = "lam y: r-``{a1}. wftrec (r,y,H)" in is_the_recfun) |
|
283 |
apply typecheck |
|
284 |
apply (unfold is_recfun_def wftrec_def) |
|
13634 | 285 |
--{*Applying the substitution: must keep the quantified assumption!*} |
13165 | 286 |
apply (rule lam_cong [OF refl]) |
287 |
apply (drule underD) |
|
288 |
apply (fold is_recfun_def) |
|
289 |
apply (rule_tac t = "%z. H(?x,z)" in subst_context) |
|
290 |
apply (rule fun_extension) |
|
291 |
apply (blast intro: is_recfun_type) |
|
292 |
apply (rule lam_type [THEN restrict_type2]) |
|
293 |
apply blast |
|
294 |
apply (blast dest: transD) |
|
295 |
apply (frule spec [THEN mp], assumption) |
|
296 |
apply (subgoal_tac "<xa,a1> : r") |
|
13784 | 297 |
apply (drule_tac x1 = xa in spec [THEN mp], assumption) |
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13165
diff
changeset
|
298 |
apply (simp add: vimage_singleton_iff |
13165 | 299 |
apply_recfun is_recfun_cut) |
300 |
apply (blast dest: transD) |
|
301 |
done |
|
302 |
||
303 |
||
13356 | 304 |
subsection{*Unfolding @{term "wftrec(r,a,H)"}*} |
13165 | 305 |
|
306 |
lemma the_recfun_cut: |
|
307 |
"[| wf(r); trans(r); <b,a>:r |] |
|
308 |
==> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)" |
|
13269 | 309 |
by (blast intro: is_recfun_cut unfold_the_recfun) |
0 | 310 |
|
13165 | 311 |
(*NOT SUITABLE FOR REWRITING: it is recursive!*) |
312 |
lemma wftrec: |
|
313 |
"[| wf(r); trans(r) |] ==> |
|
314 |
wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))" |
|
315 |
apply (unfold wftrec_def) |
|
316 |
apply (subst unfold_the_recfun [unfolded is_recfun_def]) |
|
317 |
apply (simp_all add: vimage_singleton_iff [THEN iff_sym] the_recfun_cut) |
|
318 |
done |
|
319 |
||
13634 | 320 |
|
321 |
subsubsection{*Removal of the Premise @{term "trans(r)"}*} |
|
13165 | 322 |
|
323 |
(*NOT SUITABLE FOR REWRITING: it is recursive!*) |
|
324 |
lemma wfrec: |
|
325 |
"wf(r) ==> wfrec(r,a,H) = H(a, lam x:r-``{a}. wfrec(r,x,H))" |
|
326 |
apply (unfold wfrec_def) |
|
327 |
apply (erule wf_trancl [THEN wftrec, THEN ssubst]) |
|
328 |
apply (rule trans_trancl) |
|
329 |
apply (rule vimage_pair_mono [THEN restrict_lam_eq, THEN subst_context]) |
|
330 |
apply (erule r_into_trancl) |
|
331 |
apply (rule subset_refl) |
|
332 |
done |
|
0 | 333 |
|
13165 | 334 |
(*This form avoids giant explosions in proofs. NOTE USE OF == *) |
335 |
lemma def_wfrec: |
|
336 |
"[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==> |
|
337 |
h(a) = H(a, lam x: r-``{a}. h(x))" |
|
338 |
apply simp |
|
339 |
apply (elim wfrec) |
|
340 |
done |
|
341 |
||
342 |
lemma wfrec_type: |
|
343 |
"[| wf(r); a:A; field(r)<=A; |
|
344 |
!!x u. [| x: A; u: Pi(r-``{x}, B) |] ==> H(x,u) : B(x) |
|
345 |
|] ==> wfrec(r,a,H) : B(a)" |
|
13784 | 346 |
apply (rule_tac a = a in wf_induct2, assumption+) |
13165 | 347 |
apply (subst wfrec, assumption) |
348 |
apply (simp add: lam_type underD) |
|
349 |
done |
|
350 |
||
351 |
||
352 |
lemma wfrec_on: |
|
353 |
"[| wf[A](r); a: A |] ==> |
|
354 |
wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))" |
|
355 |
apply (unfold wf_on_def wfrec_on_def) |
|
356 |
apply (erule wfrec [THEN trans]) |
|
357 |
apply (simp add: vimage_Int_square cons_subset_iff) |
|
358 |
done |
|
0 | 359 |
|
13634 | 360 |
text{*Minimal-element characterization of well-foundedness*} |
13165 | 361 |
lemma wf_eq_minimal: |
362 |
"wf(r) <-> (ALL Q x. x:Q --> (EX z:Q. ALL y. <y,z>:r --> y~:Q))" |
|
13634 | 363 |
by (unfold wf_def, blast) |
364 |
||
13165 | 365 |
|
366 |
ML |
|
367 |
{* |
|
368 |
val wf_def = thm "wf_def"; |
|
369 |
val wf_on_def = thm "wf_on_def"; |
|
0 | 370 |
|
13165 | 371 |
val wf_imp_wf_on = thm "wf_imp_wf_on"; |
372 |
val wf_on_field_imp_wf = thm "wf_on_field_imp_wf"; |
|
373 |
val wf_iff_wf_on_field = thm "wf_iff_wf_on_field"; |
|
374 |
val wf_on_subset_A = thm "wf_on_subset_A"; |
|
375 |
val wf_on_subset_r = thm "wf_on_subset_r"; |
|
376 |
val wf_onI = thm "wf_onI"; |
|
377 |
val wf_onI2 = thm "wf_onI2"; |
|
378 |
val wf_induct = thm "wf_induct"; |
|
379 |
val wf_induct2 = thm "wf_induct2"; |
|
380 |
val field_Int_square = thm "field_Int_square"; |
|
381 |
val wf_on_induct = thm "wf_on_induct"; |
|
382 |
val wfI = thm "wfI"; |
|
383 |
val wf_not_refl = thm "wf_not_refl"; |
|
384 |
val wf_not_sym = thm "wf_not_sym"; |
|
385 |
val wf_asym = thm "wf_asym"; |
|
386 |
val wf_on_not_refl = thm "wf_on_not_refl"; |
|
387 |
val wf_on_not_sym = thm "wf_on_not_sym"; |
|
388 |
val wf_on_asym = thm "wf_on_asym"; |
|
389 |
val wf_on_chain3 = thm "wf_on_chain3"; |
|
390 |
val wf_on_trancl = thm "wf_on_trancl"; |
|
391 |
val wf_trancl = thm "wf_trancl"; |
|
392 |
val underI = thm "underI"; |
|
393 |
val underD = thm "underD"; |
|
394 |
val is_recfun_type = thm "is_recfun_type"; |
|
395 |
val apply_recfun = thm "apply_recfun"; |
|
396 |
val is_recfun_equal = thm "is_recfun_equal"; |
|
397 |
val is_recfun_cut = thm "is_recfun_cut"; |
|
398 |
val is_recfun_functional = thm "is_recfun_functional"; |
|
399 |
val is_the_recfun = thm "is_the_recfun"; |
|
400 |
val unfold_the_recfun = thm "unfold_the_recfun"; |
|
401 |
val the_recfun_cut = thm "the_recfun_cut"; |
|
402 |
val wftrec = thm "wftrec"; |
|
403 |
val wfrec = thm "wfrec"; |
|
404 |
val def_wfrec = thm "def_wfrec"; |
|
405 |
val wfrec_type = thm "wfrec_type"; |
|
406 |
val wfrec_on = thm "wfrec_on"; |
|
407 |
val wf_eq_minimal = thm "wf_eq_minimal"; |
|
408 |
*} |
|
435 | 409 |
|
0 | 410 |
end |