| author | paulson | 
| Mon, 12 Jan 2004 16:45:35 +0100 | |
| changeset 14352 | a8b1a44d8264 | 
| parent 14169 | 0590de71a016 | 
| child 15131 | c69542757a4d | 
| 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 *) theory Lfp = Product_Type: constdefs lfp :: "['a set \<Rightarrow> 'a set] \<Rightarrow> 'a set" "lfp(f) == Inter({u. f(u) <= u})" (*least fixed point*) end