* Pure: method 'atomize' presents local goal premises as object-level
statements (atomic meta-level propositions); setup controlled via
rewrite rules declarations of 'atomize' attribute; example
application: 'induct' method with proper rule statements in improper
proof *scripts*;
(* Title: HOL/Lfp.thy
ID: $Id$
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1992 University of Cambridge
The Knaster-Tarski Theorem
*)
Lfp = Product_Type +
constdefs
lfp :: ['a set=>'a set] => 'a set
"lfp(f) == Inter({u. f(u) <= u})" (*least fixed point*)
end