added pretend_use;
simplified ML handling;
loaded_files: include thy;
perform Remove *before* actual deletion;
perform: made bullet proof;
(*  Title:      HOL/UNITY/SubstAx
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1998  University of Cambridge
Weak LeadsTo relation (restricted to the set of reachable states)
*)
SubstAx = WFair + Constrains + 
consts
   LeadsTo :: "['a set, 'a set] => 'a program set"       (infixl 60)
defs
   LeadsTo_def
    "A LeadsTo B == {F. F : (reachable F Int A) leadsTo  B}"
end