- Datatype package now also supports arbitrarily branching datatypes
(using function types).
- Added new simplification procedure for proving distinctness of
constructors.
- dtK is now a reference.
(* 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