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