src/HOL/Tools/Sledgehammer/sledgehammer_run.ML
author blanchet
Fri, 20 Jul 2012 22:19:46 +0200
changeset 49399 83dc102041e6
parent 49396 1b7d798460bb
child 49409 82fc8c956cdc
permissions -rw-r--r--
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_run.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_RUN =
    10 sig
    11   type minimize_command = ATP_Proof_Reconstruct.minimize_command
    12   type fact_override = Sledgehammer_Fact.fact_override
    13   type mode = Sledgehammer_Provers.mode
    14   type params = Sledgehammer_Provers.params
    15 
    16   val someN : string
    17   val noneN : string
    18   val timeoutN : string
    19   val unknownN : string
    20   val run_sledgehammer :
    21     params -> mode -> int -> fact_override
    22     -> ((string * string list) list -> string -> minimize_command)
    23     -> Proof.state -> bool * (string * Proof.state)
    24 end;
    25 
    26 structure Sledgehammer_Run : SLEDGEHAMMER_RUN =
    27 struct
    28 
    29 open ATP_Util
    30 open ATP_Problem_Generate
    31 open ATP_Proof_Reconstruct
    32 open Sledgehammer_Util
    33 open Sledgehammer_Fact
    34 open Sledgehammer_Provers
    35 open Sledgehammer_Minimize
    36 open Sledgehammer_MaSh
    37 
    38 val someN = "some"
    39 val noneN = "none"
    40 val timeoutN = "timeout"
    41 val unknownN = "unknown"
    42 
    43 val ordered_outcome_codes = [someN, unknownN, timeoutN, noneN]
    44 
    45 fun max_outcome_code codes =
    46   NONE
    47   |> fold (fn candidate =>
    48               fn accum as SOME _ => accum
    49                | NONE => if member (op =) codes candidate then SOME candidate
    50                          else NONE)
    51           ordered_outcome_codes
    52   |> the_default unknownN
    53 
    54 fun prover_description ctxt ({verbose, blocking, ...} : params) name num_facts i
    55                        n goal =
    56   (quote name,
    57    (if verbose then
    58       " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts
    59     else
    60       "") ^
    61    " on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^
    62    (if blocking then "."
    63     else "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))))
    64 
    65 fun launch_prover (params as {debug, verbose, blocking, max_facts, slice,
    66                               timeout, expect, ...})
    67         mode minimize_command only {state, goal, subgoal, subgoal_count, facts}
    68         name =
    69   let
    70     val ctxt = Proof.context_of state
    71     val hard_timeout = Time.+ (timeout, timeout)
    72     val birth_time = Time.now ()
    73     val death_time = Time.+ (birth_time, hard_timeout)
    74     val max_facts =
    75       max_facts |> the_default (default_max_facts_for_prover ctxt slice name)
    76     val num_facts = length facts |> not only ? Integer.min max_facts
    77     fun desc () =
    78       prover_description ctxt params name num_facts subgoal subgoal_count goal
    79     val problem =
    80       {state = state, goal = goal, subgoal = subgoal,
    81        subgoal_count = subgoal_count,
    82        facts = facts
    83                |> not (Sledgehammer_Provers.is_ho_atp ctxt name)
    84                   ? filter_out (curry (op =) Induction o snd o snd o fst
    85                                 o untranslated_fact)
    86                |> take num_facts}
    87     fun really_go () =
    88       problem
    89       |> get_minimizing_prover ctxt mode
    90              (mash_learn_proof ctxt params name (prop_of goal)
    91                                (map untranslated_fact facts))
    92              name params minimize_command
    93       |> (fn {outcome, preplay, message, message_tail, ...} =>
    94              (if outcome = SOME ATP_Proof.TimedOut then timeoutN
    95               else if is_some outcome then noneN
    96               else someN, fn () => message (preplay ()) ^ message_tail))
    97     fun go () =
    98       let
    99         val (outcome_code, message) =
   100           if debug then
   101             really_go ()
   102           else
   103             (really_go ()
   104              handle ERROR msg => (unknownN, fn () => "Error: " ^ msg ^ "\n")
   105                   | exn =>
   106                     if Exn.is_interrupt exn then
   107                       reraise exn
   108                     else
   109                       (unknownN, fn () => "Internal error:\n" ^
   110                                           ML_Compiler.exn_message exn ^ "\n"))
   111         val _ =
   112           (* The "expect" argument is deliberately ignored if the prover is
   113              missing so that the "Metis_Examples" can be processed on any
   114              machine. *)
   115           if expect = "" orelse outcome_code = expect orelse
   116              not (is_prover_installed ctxt name) then
   117             ()
   118           else if blocking then
   119             error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
   120           else
   121             warning ("Unexpected outcome: " ^ quote outcome_code ^ ".");
   122       in (outcome_code, message) end
   123   in
   124     if mode = Auto_Try then
   125       let val (outcome_code, message) = TimeLimit.timeLimit timeout go () in
   126         (outcome_code,
   127          state
   128          |> outcome_code = someN
   129             ? Proof.goal_message (fn () =>
   130                   [Pretty.str "",
   131                    Pretty.mark Isabelle_Markup.hilite (Pretty.str (message ()))]
   132                   |> Pretty.chunks))
   133       end
   134     else if blocking then
   135       let
   136         val (outcome_code, message) = TimeLimit.timeLimit hard_timeout go ()
   137       in
   138         (if outcome_code = someN orelse mode = Normal then
   139            quote name ^ ": " ^ message ()
   140          else
   141            "")
   142         |> Async_Manager.break_into_chunks
   143         |> List.app Output.urgent_message;
   144         (outcome_code, state)
   145       end
   146     else
   147       (Async_Manager.launch SledgehammerN birth_time death_time (desc ())
   148                             ((fn (outcome_code, message) =>
   149                                  (verbose orelse outcome_code = someN,
   150                                   message ())) o go);
   151        (unknownN, state))
   152   end
   153 
   154 fun class_of_smt_solver ctxt name =
   155   ctxt |> select_smt_solver name
   156        |> SMT_Config.solver_class_of |> SMT_Utils.string_of_class
   157 
   158 val auto_try_max_facts_divisor = 2 (* FUDGE *)
   159 
   160 fun run_sledgehammer (params as {debug, verbose, blocking, provers, max_facts,
   161                                  slice, ...})
   162         mode i (fact_override as {only, ...}) minimize_command state =
   163   if null provers then
   164     error "No prover is set."
   165   else case subgoal_count state of
   166     0 => (Output.urgent_message "No subgoal!"; (false, (noneN, state)))
   167   | n =>
   168     let
   169       val _ = Proof.assert_backward state
   170       val print = if mode = Normal then Output.urgent_message else K ()
   171       val state =
   172         state |> Proof.map_context (Config.put SMT_Config.verbose debug)
   173       val ctxt = Proof.context_of state
   174       val {facts = chained_ths, goal, ...} = Proof.goal state
   175       val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal i
   176       val ho_atp = exists (Sledgehammer_Provers.is_ho_atp ctxt) provers
   177       val reserved = reserved_isar_keyword_table ()
   178       val css_table = clasimpset_rule_table_of ctxt
   179       val facts =
   180         nearly_all_facts ctxt ho_atp fact_override reserved css_table
   181                          chained_ths hyp_ts concl_t
   182       val _ = () |> not blocking ? kill_provers
   183       val _ = case find_first (not o is_prover_supported ctxt) provers of
   184                 SOME name => error ("No such prover: " ^ name ^ ".")
   185               | NONE => ()
   186       val _ = print "Sledgehammering..."
   187       val (smts, (ueq_atps, full_atps)) =
   188         provers |> List.partition (is_smt_prover ctxt)
   189                 ||> List.partition (is_unit_equational_atp ctxt)
   190       fun launch_provers state get_facts translate provers =
   191         let
   192           val facts = get_facts ()
   193           val num_facts = length facts
   194           val facts = facts ~~ (0 upto num_facts - 1)
   195                       |> map (translate num_facts)
   196           val problem =
   197             {state = state, goal = goal, subgoal = i, subgoal_count = n,
   198              facts = facts}
   199           val launch = launch_prover params mode minimize_command only
   200         in
   201           if mode = Auto_Try orelse mode = Try then
   202             (unknownN, state)
   203             |> fold (fn prover => fn accum as (outcome_code, _) =>
   204                         if outcome_code = someN then accum
   205                         else launch problem prover)
   206                     provers
   207           else
   208             provers
   209             |> (if blocking then Par_List.map else map)
   210                    (launch problem #> fst)
   211             |> max_outcome_code |> rpair state
   212         end
   213       fun get_facts label is_appropriate_prop provers =
   214         let
   215           val max_max_facts =
   216             case max_facts of
   217               SOME n => n
   218             | NONE =>
   219               0 |> fold (Integer.max o default_max_facts_for_prover ctxt slice)
   220                         provers
   221                 |> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)
   222         in
   223           facts
   224           |> (case is_appropriate_prop of
   225                 SOME is_app => filter (is_app o prop_of o snd)
   226               | NONE => I)
   227           |> relevant_facts ctxt params (hd provers) max_max_facts fact_override
   228                             hyp_ts concl_t
   229           |> map (apfst (apfst (fn name => name ())))
   230           |> tap (fn facts =>
   231                      if debug then
   232                        label ^ plural_s (length provers) ^ ": " ^
   233                        (if null facts then
   234                           "Found no relevant facts."
   235                         else
   236                           "Including (up to) " ^ string_of_int (length facts) ^
   237                           " relevant fact" ^ plural_s (length facts) ^ ":\n" ^
   238                           (facts |> map (fst o fst) |> space_implode " ") ^ ".")
   239                        |> print
   240                      else
   241                        ())
   242         end
   243       fun launch_atps label is_appropriate_prop atps accum =
   244         if null atps then
   245           accum
   246         else if is_some is_appropriate_prop andalso
   247                 not (the is_appropriate_prop concl_t) then
   248           (if verbose orelse length atps = length provers then
   249              "Goal outside the scope of " ^
   250              space_implode " " (serial_commas "and" (map quote atps)) ^ "."
   251              |> Output.urgent_message
   252            else
   253              ();
   254            accum)
   255         else
   256           launch_provers state (get_facts label is_appropriate_prop o K atps)
   257                          (K (Untranslated_Fact o fst)) atps
   258       fun launch_smts accum =
   259         if null smts then
   260           accum
   261         else
   262           let
   263             val facts = get_facts "SMT solver" NONE smts
   264             val weight = SMT_Weighted_Fact oo weight_smt_fact ctxt
   265           in
   266             smts |> map (`(class_of_smt_solver ctxt))
   267                  |> AList.group (op =)
   268                  |> map (snd #> launch_provers state (K facts) weight #> fst)
   269                  |> max_outcome_code |> rpair state
   270           end
   271       val launch_full_atps = launch_atps "ATP" NONE full_atps
   272       val launch_ueq_atps =
   273         launch_atps "Unit equational provers" (SOME is_unit_equality) ueq_atps
   274       fun launch_atps_and_smt_solvers () =
   275         [launch_full_atps, launch_smts, launch_ueq_atps]
   276         |> Par_List.map (fn f => ignore (f (unknownN, state)))
   277         handle ERROR msg => (print ("Error: " ^ msg); error msg)
   278       fun maybe f (accum as (outcome_code, _)) =
   279         accum |> (mode = Normal orelse outcome_code <> someN) ? f
   280     in
   281       (unknownN, state)
   282       |> (if blocking then
   283             launch_full_atps
   284             #> mode <> Auto_Try ? (maybe launch_ueq_atps #> maybe launch_smts)
   285           else
   286             (fn p => Future.fork (tap launch_atps_and_smt_solvers) |> K p))
   287       handle TimeLimit.TimeOut =>
   288              (print "Sledgehammer ran out of time."; (unknownN, state))
   289     end
   290     |> `(fn (outcome_code, _) => outcome_code = someN)
   291 
   292 end;