author | clasohm |
Fri, 08 Mar 1996 13:11:09 +0100 | |
changeset 1558 | 9c6ebfab4e05 |
parent 1475 | 7f5a4cd08209 |
child 3947 | eb707467f8c5 |
permissions | -rw-r--r-- |
1475 | 1 |
(* Title: HOL/gfp.thy |
923 | 2 |
ID: $Id$ |
1475 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Greatest fixed points (requires Lfp too!) |
|
7 |
*) |
|
8 |
||
9 |
Gfp = Lfp + |
|
1558 | 10 |
|
11 |
constdefs |
|
12 |
gfp :: ['a set=>'a set] => 'a set |
|
13 |
"gfp(f) == Union({u. u <= f(u)})" (*greatest fixed point*) |
|
14 |
||
923 | 15 |
end |