src/HOL/Tools/Sledgehammer/sledgehammer.ML
author blanchet
Mon, 16 Aug 2010 16:58:45 +0200
changeset 38699 131f7c46cf2c
parent 38514 581a402a80f0
child 38740 32391240695f
permissions -rw-r--r--
more debug output
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer.ML
     2     Author:     Fabian Immler, TU Muenchen
     3     Author:     Makarius
     4     Author:     Jasmin Blanchette, TU Muenchen
     5 
     6 Sledgehammer's heart.
     7 *)
     8 
     9 signature SLEDGEHAMMER =
    10 sig
    11   type failure = ATP_Systems.failure
    12   type relevance_override = Sledgehammer_Fact_Filter.relevance_override
    13   type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command
    14   type params =
    15     {debug: bool,
    16      verbose: bool,
    17      overlord: bool,
    18      atps: string list,
    19      full_types: bool,
    20      explicit_apply: bool,
    21      relevance_threshold: real,
    22      relevance_convergence: real,
    23      theory_relevant: bool option,
    24      defs_relevant: bool,
    25      isar_proof: bool,
    26      isar_shrink_factor: int,
    27      timeout: Time.time,
    28      minimize_timeout: Time.time}
    29   type problem =
    30     {subgoal: int,
    31      goal: Proof.context * (thm list * thm),
    32      relevance_override: relevance_override,
    33      axioms: (string * thm) list option}
    34   type prover_result =
    35     {outcome: failure option,
    36      message: string,
    37      pool: string Symtab.table,
    38      used_thm_names: string list,
    39      atp_run_time_in_msecs: int,
    40      output: string,
    41      proof: string,
    42      axiom_names: string Vector.vector,
    43      conjecture_shape: int list list}
    44   type prover = params -> minimize_command -> problem -> prover_result
    45 
    46   val dest_dir : string Config.T
    47   val problem_prefix : string Config.T
    48   val measure_runtime : bool Config.T
    49   val kill_atps: unit -> unit
    50   val running_atps: unit -> unit
    51   val messages: int option -> unit
    52   val get_prover_fun : theory -> string -> prover
    53   val run_sledgehammer :
    54     params -> int -> relevance_override -> (string -> minimize_command)
    55     -> Proof.state -> unit
    56   val setup : theory -> theory
    57 end;
    58 
    59 structure Sledgehammer : SLEDGEHAMMER =
    60 struct
    61 
    62 open ATP_Problem
    63 open ATP_Systems
    64 open Metis_Clauses
    65 open Sledgehammer_Util
    66 open Sledgehammer_Fact_Filter
    67 open Sledgehammer_Translate
    68 open Sledgehammer_Proof_Reconstruct
    69 
    70 
    71 (** The Sledgehammer **)
    72 
    73 (* Identifier to distinguish Sledgehammer from other tools using
    74    "Async_Manager". *)
    75 val das_Tool = "Sledgehammer"
    76 
    77 fun kill_atps () = Async_Manager.kill_threads das_Tool "ATPs"
    78 fun running_atps () = Async_Manager.running_threads das_Tool "ATPs"
    79 val messages = Async_Manager.thread_messages das_Tool "ATP"
    80 
    81 (** problems, results, provers, etc. **)
    82 
    83 type params =
    84   {debug: bool,
    85    verbose: bool,
    86    overlord: bool,
    87    atps: string list,
    88    full_types: bool,
    89    explicit_apply: bool,
    90    relevance_threshold: real,
    91    relevance_convergence: real,
    92    theory_relevant: bool option,
    93    defs_relevant: bool,
    94    isar_proof: bool,
    95    isar_shrink_factor: int,
    96    timeout: Time.time,
    97    minimize_timeout: Time.time}
    98 
    99 type problem =
   100   {subgoal: int,
   101    goal: Proof.context * (thm list * thm),
   102    relevance_override: relevance_override,
   103    axioms: (string * thm) list option}
   104 
   105 type prover_result =
   106   {outcome: failure option,
   107    message: string,
   108    pool: string Symtab.table,
   109    used_thm_names: string list,
   110    atp_run_time_in_msecs: int,
   111    output: string,
   112    proof: string,
   113    axiom_names: string Vector.vector,
   114    conjecture_shape: int list list}
   115 
   116 type prover = params -> minimize_command -> problem -> prover_result
   117 
   118 (* configuration attributes *)
   119 
   120 val (dest_dir, dest_dir_setup) = Attrib.config_string "atp_dest_dir" (K "");
   121   (*Empty string means create files in Isabelle's temporary files directory.*)
   122 
   123 val (problem_prefix, problem_prefix_setup) =
   124   Attrib.config_string "atp_problem_prefix" (K "prob");
   125 
   126 val (measure_runtime, measure_runtime_setup) =
   127   Attrib.config_bool "atp_measure_runtime" (K false);
   128 
   129 fun with_path cleanup after f path =
   130   Exn.capture f path
   131   |> tap (fn _ => cleanup path)
   132   |> Exn.release
   133   |> tap (after path)
   134 
   135 (* Splits by the first possible of a list of delimiters. *)
   136 fun extract_proof delims output =
   137   case pairself (find_first (fn s => String.isSubstring s output))
   138                 (ListPair.unzip delims) of
   139     (SOME begin_delim, SOME end_delim) =>
   140     (output |> first_field begin_delim |> the |> snd
   141             |> first_field end_delim |> the |> fst
   142             |> first_field "\n" |> the |> snd
   143      handle Option.Option => "")
   144   | _ => ""
   145 
   146 fun extract_proof_and_outcome complete res_code proof_delims known_failures
   147                               output =
   148   case known_failure_in_output output known_failures of
   149     NONE => (case extract_proof proof_delims output of
   150              "" => ("", SOME MalformedOutput)
   151            | proof => if res_code = 0 then (proof, NONE)
   152                       else ("", SOME UnknownError))
   153   | SOME failure =>
   154     ("", SOME (if failure = IncompleteUnprovable andalso complete then
   155                  Unprovable
   156                else
   157                  failure))
   158 
   159 fun extract_clause_sequence output =
   160   let
   161     val tokens_of = String.tokens (not o Char.isAlphaNum)
   162     fun extract_num ("clause" :: (ss as _ :: _)) =
   163     Int.fromString (List.last ss)
   164       | extract_num _ = NONE
   165   in output |> split_lines |> map_filter (extract_num o tokens_of) end
   166 
   167 val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
   168 
   169 val parse_clause_formula_pair =
   170   $$ "(" |-- scan_integer --| $$ "," -- Symbol.scan_id
   171   --| Scan.repeat ($$ "," |-- Symbol.scan_id) --| $$ ")"
   172   --| Scan.option ($$ ",")
   173 val parse_clause_formula_relation =
   174   Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
   175   |-- Scan.repeat parse_clause_formula_pair
   176 val extract_clause_formula_relation =
   177   Substring.full
   178   #> Substring.position set_ClauseFormulaRelationN
   179   #> snd #> Substring.string #> strip_spaces #> explode
   180   #> parse_clause_formula_relation #> fst
   181 
   182 fun repair_conjecture_shape_and_theorem_names output conjecture_shape
   183                                               thm_names =
   184   if String.isSubstring set_ClauseFormulaRelationN output then
   185     (* This is a hack required for keeping track of axioms after they have been
   186        clausified by SPASS's Flotter tool. The "SPASS_TPTP" script is also part
   187        of this hack. *)
   188     let
   189       val j0 = hd (hd conjecture_shape)
   190       val seq = extract_clause_sequence output
   191       val name_map = extract_clause_formula_relation output
   192       fun renumber_conjecture j =
   193         AList.find (op =) name_map (conjecture_prefix ^ Int.toString (j - j0))
   194         |> map (fn s => find_index (curry (op =) s) seq + 1)
   195     in
   196       (conjecture_shape |> map (maps renumber_conjecture),
   197        seq |> map (fn j => case j |> AList.lookup (op =) name_map |> the
   198                                   |> try (unprefix axiom_prefix) of
   199                              SOME s' => undo_ascii_of s'
   200                            | NONE => "")
   201            |> Vector.fromList)
   202     end
   203   else
   204     (conjecture_shape, thm_names)
   205 
   206 
   207 (* generic TPTP-based provers *)
   208 
   209 fun prover_fun atp_name
   210         {exec, required_execs, arguments, proof_delims, known_failures,
   211          max_new_relevant_facts_per_iter, prefers_theory_relevant,
   212          explicit_forall}
   213         ({debug, verbose, overlord, full_types, explicit_apply,
   214           relevance_threshold, relevance_convergence, theory_relevant,
   215           defs_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
   216         minimize_command
   217         ({subgoal, goal, relevance_override, axioms} : problem) =
   218   let
   219     val (ctxt, (_, th)) = goal;
   220     val thy = ProofContext.theory_of ctxt
   221     val (params, hyp_ts, concl_t) = strip_subgoal th subgoal
   222 
   223     fun print s = (priority s; if debug then tracing s else ())
   224     fun print_v f = () |> verbose ? print o f
   225     fun print_d f = () |> debug ? print o f
   226 
   227     val the_axioms =
   228       case axioms of
   229         SOME axioms => axioms
   230       | NONE =>
   231         (print_d (fn () => "Selecting relevant facts for " ^ quote atp_name ^
   232                            ".");
   233          relevant_facts full_types relevance_threshold relevance_convergence
   234                         defs_relevant max_new_relevant_facts_per_iter
   235                         (the_default prefers_theory_relevant theory_relevant)
   236                         relevance_override goal hyp_ts concl_t
   237          |> tap ((fn n => print_v (fn () =>
   238                       "Selected " ^ string_of_int n ^ " fact" ^ plural_s n ^
   239                       " for " ^ quote atp_name ^ ".")) o length))
   240 
   241     (* path to unique problem file *)
   242     val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
   243                        else Config.get ctxt dest_dir;
   244     val the_problem_prefix = Config.get ctxt problem_prefix;
   245     fun prob_pathname nr =
   246       let
   247         val probfile =
   248           Path.basic ((if overlord then "prob_" ^ atp_name
   249                        else the_problem_prefix ^ serial_string ())
   250                       ^ "_" ^ string_of_int nr)
   251       in
   252         if the_dest_dir = "" then File.tmp_path probfile
   253         else if File.exists (Path.explode the_dest_dir)
   254         then Path.append (Path.explode the_dest_dir) probfile
   255         else error ("No such directory: " ^ the_dest_dir ^ ".")
   256       end;
   257 
   258     val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
   259     (* write out problem file and call prover *)
   260     fun command_line complete probfile =
   261       let
   262         val core = File.shell_path command ^ " " ^ arguments complete timeout ^
   263                    " " ^ File.shell_path probfile
   264       in
   265         (if Config.get ctxt measure_runtime then
   266            "TIMEFORMAT='%3U'; { time " ^ core ^ " ; }"
   267          else
   268            "exec " ^ core) ^ " 2>&1"
   269       end
   270     fun split_time s =
   271       let
   272         val split = String.tokens (fn c => str c = "\n");
   273         val (output, t) = s |> split |> split_last |> apfst cat_lines;
   274         fun as_num f = f >> (fst o read_int);
   275         val num = as_num (Scan.many1 Symbol.is_ascii_digit);
   276         val digit = Scan.one Symbol.is_ascii_digit;
   277         val num3 = as_num (digit ::: digit ::: (digit >> single));
   278         val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b);
   279         val as_time = the_default 0 o Scan.read Symbol.stopper time o explode;
   280       in (output, as_time t) end;
   281     fun run_on probfile =
   282       case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
   283         (home_var, _) :: _ =>
   284         error ("The environment variable " ^ quote home_var ^ " is not set.")
   285       | [] =>
   286         if File.exists command then
   287           let
   288             fun do_run complete =
   289               let
   290                 val command = command_line complete probfile
   291                 val ((output, msecs), res_code) =
   292                   bash_output command
   293                   |>> (if overlord then
   294                          prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
   295                        else
   296                          I)
   297                   |>> (if Config.get ctxt measure_runtime then split_time
   298                        else rpair 0)
   299                 val (proof, outcome) =
   300                   extract_proof_and_outcome complete res_code proof_delims
   301                                             known_failures output
   302               in (output, msecs, proof, outcome) end
   303             val _ = print_d (fn () => "Preparing problem for " ^
   304                                       quote atp_name ^ "...")
   305             val readable_names = debug andalso overlord
   306             val (problem, pool, conjecture_offset, axiom_names) =
   307               prepare_problem ctxt readable_names explicit_forall full_types
   308                               explicit_apply hyp_ts concl_t the_axioms
   309             val _ = File.write_list probfile (strings_for_tptp_problem problem)
   310             val conjecture_shape =
   311               conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
   312               |> map single
   313             val _ = print_d (fn () => "Running " ^ quote atp_name ^ "...")
   314             val result =
   315               do_run false
   316               |> (fn (_, msecs0, _, SOME _) =>
   317                      do_run true
   318                      |> (fn (output, msecs, proof, outcome) =>
   319                             (output, msecs0 + msecs, proof, outcome))
   320                    | result => result)
   321           in ((pool, conjecture_shape, axiom_names), result) end
   322         else
   323           error ("Bad executable: " ^ Path.implode command ^ ".")
   324 
   325     (* If the problem file has not been exported, remove it; otherwise, export
   326        the proof file too. *)
   327     fun cleanup probfile =
   328       if the_dest_dir = "" then try File.rm probfile else NONE
   329     fun export probfile (_, (output, _, _, _)) =
   330       if the_dest_dir = "" then
   331         ()
   332       else
   333         File.write (Path.explode (Path.implode probfile ^ "_proof")) output
   334 
   335     val ((pool, conjecture_shape, axiom_names),
   336          (output, msecs, proof, outcome)) =
   337       with_path cleanup export run_on (prob_pathname subgoal)
   338     val (conjecture_shape, axiom_names) =
   339       repair_conjecture_shape_and_theorem_names output conjecture_shape
   340                                                 axiom_names
   341 
   342     val (message, used_thm_names) =
   343       case outcome of
   344         NONE =>
   345         proof_text isar_proof
   346             (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
   347             (full_types, minimize_command, proof, axiom_names, th, subgoal)
   348       | SOME failure => (string_for_failure failure ^ "\n", [])
   349   in
   350     {outcome = outcome, message = message, pool = pool,
   351      used_thm_names = used_thm_names, atp_run_time_in_msecs = msecs,
   352      output = output, proof = proof, axiom_names = axiom_names,
   353      conjecture_shape = conjecture_shape}
   354   end
   355 
   356 fun get_prover_fun thy name = prover_fun name (get_prover thy name)
   357 
   358 fun start_prover_thread (params as {verbose, full_types, timeout, ...}) i n
   359                         relevance_override minimize_command proof_state
   360                         atp_name =
   361   let
   362     val thy = Proof.theory_of proof_state
   363     val birth_time = Time.now ()
   364     val death_time = Time.+ (birth_time, timeout)
   365     val prover = get_prover_fun thy atp_name
   366     val {context = ctxt, facts, goal} = Proof.goal proof_state;
   367     val desc =
   368       "ATP " ^ quote atp_name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^
   369       Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i));
   370   in
   371     Async_Manager.launch das_Tool verbose birth_time death_time desc
   372         (fn () =>
   373             let
   374               val problem =
   375                 {subgoal = i, goal = (ctxt, (facts, goal)),
   376                  relevance_override = relevance_override, axioms = NONE}
   377             in
   378               prover params (minimize_command atp_name) problem |> #message
   379               handle ERROR message => "Error: " ^ message ^ "\n"
   380                    | exn => "Internal error: \n" ^ ML_Compiler.exn_message exn ^
   381                             "\n"
   382             end)
   383   end
   384 
   385 fun run_sledgehammer {atps = [], ...} _ _ _ _ = error "No ATP is set."
   386   | run_sledgehammer (params as {atps, ...}) i relevance_override
   387                      minimize_command state =
   388     case subgoal_count state of
   389       0 => priority "No subgoal!"
   390     | n =>
   391       let
   392         val _ = kill_atps ()
   393         val _ = priority "Sledgehammering..."
   394         val _ = app (start_prover_thread params i n relevance_override
   395                                          minimize_command state) atps
   396       in () end
   397 
   398 val setup =
   399   dest_dir_setup
   400   #> problem_prefix_setup
   401   #> measure_runtime_setup
   402 
   403 end;