| author | clasohm | 
| Wed, 04 Oct 1995 13:12:14 +0100 | |
| changeset 1266 | 3ae9fe3c0f68 | 
| parent 923 | ff1574a81019 | 
| child 1370 | 7361ac9b024d | 
| permissions | -rw-r--r-- | 
(* Title: HOL/gfp.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Greatest fixed points (requires Lfp too!) *) Gfp = Lfp + consts gfp :: "['a set=>'a set] => 'a set" defs (*greatest fixed point*) gfp_def "gfp(f) == Union({u. u <= f(u)})" end