| author | paulson | 
| Mon, 15 Jul 1996 14:58:28 +0200 | |
| changeset 1862 | 74d4ae2f6fc3 | 
| parent 1558 | 9c6ebfab4e05 | 
| 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  |