src/HOL/Tools/res_atp.ML
author paulson
Wed, 05 Oct 2005 11:18:06 +0200
changeset 17764 fde495b9e24b
parent 17746 af59c748371d
child 17772 818cec5f82a4
permissions -rw-r--r--
improved process handling. tidied
     1 (*  Author: Jia Meng, Cambridge University Computer Laboratory
     2     ID: $Id$
     3     Copyright 2004 University of Cambridge
     4 
     5 ATPs with TPTP format input.
     6 *)
     7 
     8 signature RES_ATP =
     9 sig
    10   val prover: string ref
    11   val custom_spass: string list ref
    12   val destdir: string ref
    13   val hook_count: int ref
    14   val problem_name: string ref
    15   val time_limit: int ref
    16 end;
    17 
    18 structure ResAtp: RES_ATP =
    19 struct
    20 
    21 val call_atp = ref false;
    22 val hook_count = ref 0;
    23 val time_limit = ref 60;
    24 
    25 val prover = ref "E";   (* use E as the default prover *)
    26 val custom_spass =   (*specialized options for SPASS*)
    27       ref ["-Auto=0","-FullRed=0","-IORe","-IOFc","-RTaut","-RFSub","-RBSub"];
    28 
    29 val destdir = ref "";   (*Empty means write files to /tmp*)
    30 val problem_name = ref "prob";
    31 
    32 fun probfile_nosuffix _ = 
    33   if !destdir = "" then File.platform_path (File.tmp_path (Path.basic (!problem_name)))
    34   else if File.exists (File.unpack_platform_path (!destdir))
    35   then !destdir ^ "/" ^ !problem_name
    36   else error ("No such directory: " ^ !destdir);
    37 
    38 fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n;
    39 
    40 
    41 (**** For running in Isar ****)
    42 
    43 (* same function as that in res_axioms.ML *)
    44 fun repeat_RS thm1 thm2 =
    45     let val thm1' =  thm1 RS thm2 handle THM _ => thm1
    46     in
    47         if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
    48     end;
    49 
    50 (* a special version of repeat_RS *)
    51 fun repeat_someI_ex th = repeat_RS th someI_ex;
    52 
    53 fun writeln_strs _   []      = ()
    54   | writeln_strs out (s::ss) = (TextIO.output (out, s); TextIO.output (out, "\n"); writeln_strs out ss);
    55 
    56 (* write out a subgoal as tptp clauses to the file "xxxx_N"*)
    57 fun tptp_inputs_tfrees thms pf n (axclauses,classrel_clauses,arity_clauses) =
    58     let
    59       val clss = map (ResClause.make_conjecture_clause_thm) thms
    60       val (tptp_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2tptp clss)
    61       val tfree_clss = map ResClause.tfree_clause (ResLib.flat_noDup tfree_litss)
    62       val classrel_cls = map ResClause.tptp_classrelClause classrel_clauses
    63       val arity_cls = map ResClause.tptp_arity_clause arity_clauses
    64       val out = TextIO.openOut(pf n)
    65     in
    66       writeln_strs out (List.concat (map ResClause.tptp_clause axclauses));
    67       writeln_strs out (tfree_clss @ tptp_clss @ classrel_cls @ arity_cls);
    68       TextIO.closeOut out
    69     end;
    70 
    71 (* write out a subgoal in DFG format to the file "xxxx_N"*)
    72 fun dfg_inputs_tfrees thms pf n (axclauses,classrel_clauses,arity_clauses) = 
    73     let val clss = map (ResClause.make_conjecture_clause_thm) thms
    74         (*FIXME: classrel_clauses and arity_clauses*)
    75         val probN = ResClause.clauses2dfg clss (!problem_name ^ "_" ^ Int.toString n)
    76                         axclauses [] [] []    
    77 	val out = TextIO.openOut(pf n)
    78     in
    79 	writeln_strs out [probN]; TextIO.closeOut out
    80     end;
    81 
    82 
    83 (*********************************************************************)
    84 (* call prover with settings and problem file for the current subgoal *)
    85 (*********************************************************************)
    86 (* now passing in list of skolemized thms and list of sgterms to go with them *)
    87 fun watcher_call_provers sign sg_terms (childin, childout, pid) =
    88   let
    89     fun make_atp_list [] n = []
    90       | make_atp_list (sg_term::xs) n =
    91           let
    92             val goalstring = Sign.string_of_term sign sg_term
    93             val probfile = prob_pathname n
    94             val time = Int.toString (!time_limit)
    95           in
    96             debug ("goalstring in make_atp_lists is\n" ^ goalstring);
    97             debug ("problem file in watcher_call_provers is " ^ probfile);
    98             (*Avoid command arguments containing spaces: Poly/ML and SML/NJ
    99               versions of Unix.execute treat them differently!*)
   100             (*options are separated by Watcher.setting_sep, currently #"%"*)
   101             if !prover = "spass"
   102             then
   103               let val optionline = 
   104 		      if !AtpCommunication.reconstruct 
   105 		          (*Proof reconstruction works for only a limited set of 
   106 		            inference rules*)
   107                       then space_implode "%" (!custom_spass) ^
   108                            "%-DocProof%-TimeLimit=" ^ time
   109                       else "-DocProof%-SOS%-FullRed=0%-TimeLimit=" ^ time (*Auto mode*)
   110                   val _ = debug ("SPASS option string is " ^ optionline)
   111                   val _ = ResLib.helper_path "SPASS_HOME" "SPASS"
   112                     (*We've checked that SPASS is there for ATP/spassshell to run.*)
   113               in 
   114                   ([("spass", goalstring,
   115                      getenv "ISABELLE_HOME" ^ "/src/HOL/Tools/ATP/spassshell",
   116                      optionline, probfile)] @ 
   117                   (make_atp_list xs (n+1)))
   118               end
   119             else if !prover = "vampire"
   120 	    then 
   121               let val vampire = ResLib.helper_path "VAMPIRE_HOME" "vampire"
   122               in
   123                 ([("vampire", goalstring, vampire, "-m 100000%-t " ^ time, probfile)] @
   124                  (make_atp_list xs (n+1)))       (*BEWARE! spaces in options!*)
   125               end
   126       	     else if !prover = "E"
   127       	     then
   128 	       let val Eprover = ResLib.helper_path "E_HOME" "eproof"
   129 	       in
   130 		  ([("E", goalstring, Eprover, 
   131 		     "--tptp-in%-l5%-xAuto%-tAuto%--cpu-limit=" ^ time,
   132 		     probfile)] @
   133 		   (make_atp_list xs (n+1)))
   134 	       end
   135 	     else error ("Invalid prover name: " ^ !prover)
   136           end
   137 
   138     val atp_list = make_atp_list sg_terms 1
   139   in
   140     Watcher.callResProvers(childout,atp_list);
   141     debug "Sent commands to watcher!"
   142   end
   143 
   144 (*We write out problem files for each subgoal. Argument pf generates filenames,
   145   and allows the suppression of the suffix "_1" in problem-generation mode.*)
   146 fun write_problem_files pf (ctxt,th)  =
   147   let val prems = Thm.prems_of th
   148       val (clause_arr, axclauses) = ResClasimp.get_clasimp_lemmas ctxt (hd prems) 
   149               (*FIXME: hack!! need to consider relevance for all prems*)
   150       val _ = debug ("claset and simprules total clauses = " ^ 
   151                      Int.toString (Array.length clause_arr))
   152       val thy = ProofContext.theory_of ctxt
   153       val classrel_clauses = ResTypesSorts.classrel_clauses_thy thy
   154       val _ = debug ("classrel clauses = " ^ Int.toString (length classrel_clauses))
   155       val arity_clauses = ResTypesSorts.arity_clause_thy thy
   156       val _ = debug ("arity clauses = " ^ Int.toString (length arity_clauses))
   157       val write = if !prover = "spass" then dfg_inputs_tfrees else tptp_inputs_tfrees
   158       fun writenext n =
   159 	if n=0 then ()
   160 	 else
   161 	   (SELECT_GOAL
   162 	    (EVERY1 [rtac ccontr, ObjectLogic.atomize_tac, skolemize_tac, 
   163 	      METAHYPS(fn negs => 
   164 		(write (make_clauses negs) pf n 
   165 		       (axclauses,classrel_clauses,arity_clauses);
   166 		 writenext (n-1); 
   167 		 all_tac))]) n th;
   168 	    ())
   169       in writenext (length prems); clause_arr end;
   170 
   171 val last_watcher_pid =
   172   ref (NONE : (TextIO.instream * TextIO.outstream * Posix.Process.pid) option);
   173 
   174 
   175 (*writes out the current clasimpset to a tptp file;
   176   turns off xsymbol at start of function, restoring it at end    *)
   177 val isar_atp = setmp print_mode [] 
   178  (fn (ctxt, th) =>
   179   if Thm.no_prems th then ()
   180   else
   181     let
   182       val _ = (*Set up signal handler to tidy away dead processes*)
   183 	      IsaSignal.signal (IsaSignal.chld, 
   184 	        IsaSignal.SIG_HANDLE (fn _ =>
   185 		  (Posix.Process.waitpid_nh(Posix.Process.W_ANY_CHILD, []);
   186 		   debug "Child signal received\n")));  
   187       val _ = (case !last_watcher_pid of NONE => ()
   188                | SOME (_, childout, pid) => 
   189                   (debug ("Killing old watcher, pid = " ^ 
   190                           Int.toString (ResLib.intOfPid pid));
   191                    Watcher.killWatcher pid))
   192               handle OS.SysErr _ => debug "Attempt to kill watcher failed";
   193       val clause_arr = write_problem_files prob_pathname (ctxt,th)
   194       val (childin, childout, pid) = Watcher.createWatcher (th, clause_arr)
   195     in
   196       last_watcher_pid := SOME (childin, childout, pid);
   197       debug ("proof state: " ^ string_of_thm th);
   198       debug ("pid: " ^ Int.toString (ResLib.intOfPid pid));
   199       watcher_call_provers (sign_of_thm th) (Thm.prems_of th) (childin, childout, pid)
   200     end);
   201 
   202 val isar_atp_writeonly = setmp print_mode [] 
   203       (fn (ctxt,th) =>
   204        if Thm.no_prems th then ()
   205        else 
   206          let val pf = if Thm.nprems_of th = 1 then probfile_nosuffix 
   207          	      else prob_pathname
   208          in ignore (write_problem_files pf (ctxt,th)) end);
   209 
   210 
   211 (** the Isar toplevel hook **)
   212 
   213 val invoke_atp = Toplevel.unknown_proof o Toplevel.keep (fn state =>
   214   let
   215     val proof = Toplevel.proof_of state
   216     val (ctxt, (_, goal)) = Proof.get_goal proof
   217         handle Proof.STATE _ => error "No goal present";
   218     val thy = ProofContext.theory_of ctxt;
   219   in
   220     debug ("initial thm in isar_atp:\n" ^ 
   221            Pretty.string_of (ProofContext.pretty_thm ctxt goal));
   222     debug ("subgoals in isar_atp:\n" ^ 
   223            Pretty.string_of (ProofContext.pretty_term ctxt
   224              (Logic.mk_conjunction_list (Thm.prems_of goal))));
   225     debug ("number of subgoals in isar_atp: " ^ Int.toString (Thm.nprems_of goal));
   226     debug ("current theory: " ^ Context.theory_name thy);
   227     hook_count := !hook_count +1;
   228     debug ("in hook for time: " ^ Int.toString (!hook_count));
   229     ResClause.init thy;
   230     if !destdir = "" andalso !time_limit > 0 then isar_atp (ctxt, goal)
   231     else isar_atp_writeonly (ctxt, goal)
   232   end);
   233 
   234 val call_atpP =
   235   OuterSyntax.command 
   236     "ProofGeneral.call_atp" 
   237     "call automatic theorem provers" 
   238     OuterKeyword.diag
   239     (Scan.succeed (Toplevel.no_timing o invoke_atp));
   240 
   241 val _ = OuterSyntax.add_parsers [call_atpP];
   242 
   243 end;