Now Datatype.occs_in_prems prints the necessary warning ITSELF.
It is also easier to invoke and even works if the induction variable
is a parameter (rather than a free variable).
(*  Title: 	HOLCF/holcfb.thy
    ID:         $Id$
    Author: 	Franz Regensburger
    Copyright   1993  Technische Universitaet Muenchen
Basic definitions for the embedding of LCF into HOL.
*)
Holcfb = Nat + 
consts
theleast     :: "(nat=>bool)=>nat"
rules
theleast_def    "theleast(P) == (@z.(P(z) & (!n.P(n)-->z<=n)))"
end