| author | wenzelm | 
| Mon, 12 Jul 1999 22:23:31 +0200 | |
| changeset 6977 | 4781c0673e83 | 
| parent 3947 | eb707467f8c5 | 
| child 8882 | 9df44a4f1bf7 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/Lfp.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge The Knaster-Tarski Theorem *) Lfp = mono + Prod + global constdefs lfp :: ['a set=>'a set] => 'a set "lfp(f) == Inter({u. f(u) <= u})" (*least fixed point*) local end