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