10213
|
1 |
(* Title: HOL/Wellfounded_Recursion.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow
|
|
4 |
Copyright 1992 University of Cambridge
|
|
5 |
|
|
6 |
Well-founded Recursion
|
|
7 |
*)
|
|
8 |
|
|
9 |
Wellfounded_Recursion = Transitive_Closure +
|
|
10 |
|
|
11 |
constdefs
|
|
12 |
wf :: "('a * 'a)set => bool"
|
|
13 |
"wf(r) == (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))"
|
|
14 |
|
|
15 |
acyclic :: "('a*'a)set => bool"
|
|
16 |
"acyclic r == !x. (x,x) ~: r^+"
|
|
17 |
|
|
18 |
cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
|
|
19 |
"cut f r x == (%y. if (y,x):r then f y else arbitrary)"
|
|
20 |
|
|
21 |
is_recfun :: "('a * 'a)set => (('a=>'b) => ('a=>'b)) =>'a=>('a=>'b) => bool"
|
|
22 |
"is_recfun r H a f == (f = cut (%x. H (cut f r x) x) r a)"
|
|
23 |
|
|
24 |
the_recfun :: "('a * 'a)set => (('a=>'b) => ('a=>'b)) => 'a => 'a => 'b"
|
|
25 |
"the_recfun r H a == (@f. is_recfun r H a f)"
|
|
26 |
|
|
27 |
wfrec :: "('a * 'a)set => (('a=>'b) => ('a=>'b)) => 'a => 'b"
|
|
28 |
"wfrec r H == (%x. H (cut (the_recfun (trancl r) (%f v. H (cut f r v) v) x)
|
|
29 |
r x) x)"
|
|
30 |
|
|
31 |
end
|