cleaned signature;
added instantiate': ctyp option list -> cterm option list -> thm -> thm;
(* Title: HOLCF/One.thy
ID: $Id$
Author: Oscar Slotosch
Copyright 1997 Technische Universitaet Muenchen
*)
One = Lift +
types one = unit lift
constdefs
ONE :: "one"
"ONE == Def ()"
translations
"one" <= (type) "unit lift"
end