author | wenzelm |
Mon, 20 Mar 2006 21:29:04 +0100 | |
changeset 19296 | ad96f1095dca |
parent 17456 | bcf7544875b2 |
child 20140 | 98acc6d0fab6 |
permissions | -rw-r--r-- |
(* Title: CCL/Lfp.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) header {* The Knaster-Tarski Theorem *} theory Lfp imports Set uses "subset.ML" "equalities.ML" "mono.ML" begin constdefs lfp :: "['a set=>'a set] => 'a set" (*least fixed point*) "lfp(f) == Inter({u. f(u) <= u})" ML {* use_legacy_bindings (the_context ()) *} end