src/HOL/Tools/Sledgehammer/sledgehammer_run.ML
author blanchet
Fri, 27 May 2011 10:30:08 +0200
changeset 43868 b10775a669d1
parent 43862 5910dd009d0e
child 43878 ade5c84f860f
permissions -rw-r--r--
make Sledgehammer a little bit less verbose in "try"
blanchet@41335
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_run.ML
wenzelm@28477
     2
    Author:     Fabian Immler, TU Muenchen
wenzelm@32996
     3
    Author:     Makarius
blanchet@35967
     4
    Author:     Jasmin Blanchette, TU Muenchen
wenzelm@28477
     5
blanchet@38255
     6
Sledgehammer's heart.
wenzelm@28477
     7
*)
wenzelm@28477
     8
blanchet@41335
     9
signature SLEDGEHAMMER_RUN =
wenzelm@28477
    10
sig
blanchet@39232
    11
  type relevance_override = Sledgehammer_Filter.relevance_override
blanchet@40249
    12
  type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
blanchet@43862
    13
  type mode = Sledgehammer_Provers.mode
blanchet@41335
    14
  type params = Sledgehammer_Provers.params
blanchet@41511
    15
  type prover = Sledgehammer_Provers.prover
blanchet@39733
    16
blanchet@43861
    17
  val someN : string
blanchet@43861
    18
  val noneN : string
blanchet@43861
    19
  val timeoutN : string
blanchet@43861
    20
  val unknownN : string
blanchet@43517
    21
  val auto_minimize_min_facts : int Config.T
blanchet@43862
    22
  val get_minimizing_prover : Proof.context -> mode -> string -> prover
blanchet@38290
    23
  val run_sledgehammer :
blanchet@43862
    24
    params -> mode -> int -> relevance_override -> (string -> minimize_command)
blanchet@43861
    25
    -> Proof.state -> bool * (string * Proof.state)
wenzelm@28477
    26
end;
wenzelm@28477
    27
blanchet@41335
    28
structure Sledgehammer_Run : SLEDGEHAMMER_RUN =
wenzelm@28477
    29
struct
wenzelm@28477
    30
blanchet@38257
    31
open Sledgehammer_Util
blanchet@39232
    32
open Sledgehammer_Filter
blanchet@40249
    33
open Sledgehammer_ATP_Translate
blanchet@41335
    34
open Sledgehammer_Provers
blanchet@41339
    35
open Sledgehammer_Minimize
blanchet@40253
    36
blanchet@43861
    37
val someN = "some"
blanchet@43861
    38
val noneN = "none"
blanchet@43861
    39
val timeoutN = "timeout"
blanchet@43861
    40
val unknownN = "unknown"
blanchet@43861
    41
blanchet@43861
    42
val ordered_outcome_codes = [someN, unknownN, timeoutN, noneN]
blanchet@43861
    43
blanchet@43861
    44
fun max_outcome_code codes =
blanchet@43861
    45
  NONE
blanchet@43861
    46
  |> fold (fn candidate =>
blanchet@43861
    47
              fn accum as SOME _ => accum
blanchet@43861
    48
               | NONE => if member (op =) codes candidate then SOME candidate
blanchet@43861
    49
                         else NONE)
blanchet@43861
    50
          ordered_outcome_codes
blanchet@43861
    51
  |> the_default unknownN
blanchet@43861
    52
blanchet@41456
    53
fun prover_description ctxt ({verbose, blocking, ...} : params) name num_facts i
blanchet@41337
    54
                       n goal =
blanchet@43846
    55
  (name,
blanchet@43846
    56
   (if verbose then
blanchet@43846
    57
      " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts
blanchet@43846
    58
    else
blanchet@43846
    59
      "") ^
blanchet@43846
    60
   " on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^
blanchet@43846
    61
   (if blocking then
blanchet@43846
    62
      "."
blanchet@43846
    63
    else
blanchet@43846
    64
      ":\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))))
blanchet@41337
    65
blanchet@43517
    66
val auto_minimize_min_facts =
blanchet@43517
    67
  Attrib.setup_config_int @{binding sledgehammer_auto_minimize_min_facts}
blanchet@43809
    68
      (fn generic => Config.get_generic generic binary_min_facts)
blanchet@41339
    69
blanchet@43862
    70
fun get_minimizing_prover ctxt mode name
blanchet@42613
    71
        (params as {debug, verbose, explicit_apply, ...}) minimize_command
blanchet@41511
    72
        (problem as {state, subgoal, subgoal_count, facts, ...}) =
blanchet@43862
    73
  get_prover ctxt mode name params minimize_command problem
blanchet@41511
    74
  |> (fn result as {outcome, used_facts, run_time_in_msecs, message} =>
blanchet@41511
    75
         if is_some outcome then
blanchet@41511
    76
           result
blanchet@41511
    77
         else
blanchet@41511
    78
           let
blanchet@41511
    79
             val (used_facts, message) =
blanchet@43517
    80
               if length used_facts
blanchet@43517
    81
                  >= Config.get ctxt auto_minimize_min_facts then
blanchet@42613
    82
                 minimize_facts name params (SOME explicit_apply) (not verbose)
blanchet@42613
    83
                     subgoal subgoal_count state
blanchet@41511
    84
                     (filter_used_facts used_facts
blanchet@41511
    85
                          (map (apsnd single o untranslated_fact) facts))
blanchet@41511
    86
                 |>> Option.map (map fst)
blanchet@41511
    87
               else
blanchet@41511
    88
                 (SOME used_facts, message)
blanchet@41511
    89
           in
blanchet@41511
    90
             case used_facts of
blanchet@41511
    91
               SOME used_facts =>
blanchet@41511
    92
               (if debug andalso not (null used_facts) then
blanchet@41511
    93
                  facts ~~ (0 upto length facts - 1)
blanchet@41511
    94
                  |> map (fn (fact, j) =>
blanchet@41511
    95
                             fact |> untranslated_fact |> apsnd (K j))
blanchet@41511
    96
                  |> filter_used_facts used_facts
blanchet@41511
    97
                  |> map (fn ((name, _), j) => name ^ "@" ^ string_of_int j)
blanchet@41511
    98
                  |> commas
blanchet@41511
    99
                  |> enclose ("Fact" ^ plural_s (length facts) ^ " in " ^
blanchet@41511
   100
                              quote name ^ " proof (of " ^
blanchet@41511
   101
                              string_of_int (length facts) ^ "): ") "."
blanchet@41511
   102
                  |> Output.urgent_message
blanchet@41511
   103
                else
blanchet@41511
   104
                  ();
blanchet@41511
   105
                {outcome = NONE, used_facts = used_facts,
blanchet@41511
   106
                 run_time_in_msecs = run_time_in_msecs, message = message})
blanchet@41511
   107
             | NONE => result
blanchet@41511
   108
           end)
blanchet@41510
   109
blanchet@43314
   110
fun launch_prover (params as {debug, blocking, max_relevant, slicing, timeout,
blanchet@43314
   111
                              expect, ...})
blanchet@43862
   112
        mode minimize_command only
blanchet@42612
   113
        {state, goal, subgoal, subgoal_count, facts, smt_filter} name =
blanchet@41337
   114
  let
blanchet@41337
   115
    val ctxt = Proof.context_of state
blanchet@43719
   116
    val hard_timeout = Time.+ (timeout, timeout)
blanchet@41337
   117
    val birth_time = Time.now ()
blanchet@43719
   118
    val death_time = Time.+ (birth_time, hard_timeout)
blanchet@41337
   119
    val max_relevant =
blanchet@43314
   120
      max_relevant
blanchet@43314
   121
      |> the_default (default_max_relevant_for_prover ctxt slicing name)
blanchet@41337
   122
    val num_facts = length facts |> not only ? Integer.min max_relevant
blanchet@43847
   123
    fun desc () =
blanchet@41337
   124
      prover_description ctxt params name num_facts subgoal subgoal_count goal
blanchet@41337
   125
    val problem =
blanchet@41337
   126
      {state = state, goal = goal, subgoal = subgoal,
blanchet@43509
   127
       subgoal_count = subgoal_count, facts = facts |> take num_facts,
blanchet@42612
   128
       smt_filter = smt_filter}
blanchet@41501
   129
    fun really_go () =
blanchet@41511
   130
      problem
blanchet@43862
   131
      |> get_minimizing_prover ctxt mode name params (minimize_command name)
blanchet@41510
   132
      |> (fn {outcome, message, ...} =>
blanchet@43846
   133
             (if outcome = SOME ATP_Proof.TimedOut then timeoutN
blanchet@43846
   134
              else if is_some outcome then noneN
blanchet@43846
   135
              else someN, message))
blanchet@41337
   136
    fun go () =
blanchet@41337
   137
      let
blanchet@41337
   138
        val (outcome_code, message) =
blanchet@41337
   139
          if debug then
blanchet@41337
   140
            really_go ()
blanchet@41337
   141
          else
blanchet@41337
   142
            (really_go ()
blanchet@43846
   143
             handle ERROR message => (unknownN, "Error: " ^ message ^ "\n")
blanchet@41337
   144
                  | exn =>
blanchet@41337
   145
                    if Exn.is_interrupt exn then
blanchet@41337
   146
                      reraise exn
blanchet@41337
   147
                    else
blanchet@43846
   148
                      (unknownN, "Internal error:\n" ^
blanchet@43846
   149
                                ML_Compiler.exn_message exn ^ "\n"))
blanchet@41337
   150
        val _ =
blanchet@41390
   151
          (* The "expect" argument is deliberately ignored if the prover is
blanchet@41390
   152
             missing so that the "Metis_Examples" can be processed on any
blanchet@41390
   153
             machine. *)
blanchet@41390
   154
          if expect = "" orelse outcome_code = expect orelse
blanchet@41390
   155
             not (is_prover_installed ctxt name) then
blanchet@41337
   156
            ()
blanchet@41337
   157
          else if blocking then
blanchet@41337
   158
            error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
blanchet@41337
   159
          else
blanchet@41337
   160
            warning ("Unexpected outcome: " ^ quote outcome_code ^ ".");
blanchet@43846
   161
      in (outcome_code, message) end
blanchet@41337
   162
  in
blanchet@43862
   163
    if mode = Auto_Try then
blanchet@43847
   164
      let val (outcome_code, message) = TimeLimit.timeLimit timeout go () in
blanchet@43847
   165
        (outcome_code,
blanchet@43847
   166
         state
blanchet@43847
   167
         |> outcome_code = someN
blanchet@43847
   168
            ? Proof.goal_message (fn () =>
blanchet@43847
   169
                  [Pretty.str "",
blanchet@43847
   170
                   Pretty.mark Markup.hilite (Pretty.str message)]
blanchet@43847
   171
                  |> Pretty.chunks))
blanchet@41337
   172
      end
blanchet@41337
   173
    else if blocking then
blanchet@43847
   174
      let
blanchet@43847
   175
        val (outcome_code, message) = TimeLimit.timeLimit hard_timeout go ()
blanchet@43847
   176
      in
blanchet@43868
   177
        (if outcome_code = someN then message
blanchet@43868
   178
         else if mode = Normal then quote name ^ ": " ^ message
blanchet@43868
   179
         else "")
blanchet@43846
   180
        |> Async_Manager.break_into_chunks
blanchet@43846
   181
        |> List.app Output.urgent_message;
blanchet@43847
   182
        (outcome_code, state)
blanchet@41337
   183
      end
blanchet@41337
   184
    else
blanchet@43847
   185
      (Async_Manager.launch das_tool birth_time death_time (desc ())
blanchet@43846
   186
                            (apfst (curry (op <>) timeoutN) o go);
blanchet@43847
   187
       (unknownN, state))
blanchet@41337
   188
  end
blanchet@41337
   189
blanchet@41483
   190
fun class_of_smt_solver ctxt name =
blanchet@41483
   191
  ctxt |> select_smt_solver name
blanchet@41483
   192
       |> SMT_Config.solver_class_of |> SMT_Utils.string_of_class
blanchet@41483
   193
blanchet@41483
   194
(* Makes backtraces more transparent and might be more efficient as well. *)
blanchet@41483
   195
fun smart_par_list_map _ [] = []
blanchet@41483
   196
  | smart_par_list_map f [x] = [f x]
blanchet@41483
   197
  | smart_par_list_map f xs = Par_List.map f xs
blanchet@41483
   198
blanchet@41502
   199
fun dest_SMT_Weighted_Fact (SMT_Weighted_Fact p) = p
blanchet@41502
   200
  | dest_SMT_Weighted_Fact _ = raise Fail "dest_SMT_Weighted_Fact"
blanchet@41502
   201
blanchet@43862
   202
val auto_try_max_relevant_divisor = 2 (* FUDGE *)
blanchet@40241
   203
blanchet@43787
   204
fun run_sledgehammer (params as {debug, verbose, blocking, provers,
blanchet@43787
   205
                                 relevance_thresholds, max_relevant, slicing,
blanchet@43787
   206
                                 timeout, ...})
blanchet@43862
   207
        mode i (relevance_override as {only, ...}) minimize_command state =
blanchet@40240
   208
  if null provers then
blanchet@40240
   209
    error "No prover is set."
blanchet@39564
   210
  else case subgoal_count state of
blanchet@43861
   211
    0 => (Output.urgent_message "No subgoal!"; (false, (noneN, state)))
blanchet@39564
   212
  | n =>
blanchet@39564
   213
    let
blanchet@39610
   214
      val _ = Proof.assert_backward state
blanchet@43862
   215
      val print = if mode = Normal then Output.urgent_message else K ()
blanchet@41483
   216
      val state =
blanchet@41483
   217
        state |> Proof.map_context (Config.put SMT_Config.verbose debug)
blanchet@40441
   218
      val ctxt = Proof.context_of state
blanchet@40441
   219
      val {facts = chained_ths, goal, ...} = Proof.goal state
blanchet@43845
   220
      val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal i
blanchet@40240
   221
      val _ = () |> not blocking ? kill_provers
blanchet@42591
   222
      val _ = case find_first (not o is_prover_supported ctxt) provers of
blanchet@41189
   223
                SOME name => error ("No such prover: " ^ name ^ ".")
blanchet@41189
   224
              | NONE => ()
blanchet@42644
   225
      val _ = print "Sledgehammering..."
blanchet@43785
   226
      val (smts, (ueq_atps, full_atps)) =
blanchet@43785
   227
        provers |> List.partition (is_smt_prover ctxt)
blanchet@43785
   228
                ||> List.partition (is_unit_equational_atp ctxt)
blanchet@42612
   229
      fun launch_provers state get_facts translate maybe_smt_filter provers =
blanchet@41502
   230
        let
blanchet@41502
   231
          val facts = get_facts ()
blanchet@41502
   232
          val num_facts = length facts
blanchet@41502
   233
          val facts = facts ~~ (0 upto num_facts - 1)
blanchet@41502
   234
                      |> map (translate num_facts)
blanchet@41502
   235
          val problem =
blanchet@41502
   236
            {state = state, goal = goal, subgoal = i, subgoal_count = n,
blanchet@41502
   237
             facts = facts,
blanchet@42612
   238
             smt_filter = maybe_smt_filter
blanchet@41502
   239
                  (fn () => map_filter (try dest_SMT_Weighted_Fact) facts) i}
blanchet@43862
   240
          val launch = launch_prover params mode minimize_command only
blanchet@41502
   241
        in
blanchet@43862
   242
          if mode = Auto_Try orelse mode = Try then
blanchet@43861
   243
            (unknownN, state)
blanchet@43862
   244
            |> fold (fn prover => fn accum as (outcome_code, _) =>
blanchet@43861
   245
                        if outcome_code = someN then accum
blanchet@43861
   246
                        else launch problem prover)
blanchet@43861
   247
                    provers
blanchet@41502
   248
          else
blanchet@41502
   249
            provers
blanchet@43861
   250
            |> (if blocking then smart_par_list_map else map)
blanchet@43861
   251
                   (launch problem #> fst)
blanchet@43861
   252
            |> max_outcome_code |> rpair state
blanchet@41502
   253
        end
blanchet@43793
   254
      fun get_facts label is_appropriate_prop relevance_fudge provers =
blanchet@41483
   255
        let
blanchet@41483
   256
          val max_max_relevant =
blanchet@41483
   257
            case max_relevant of
blanchet@41483
   258
              SOME n => n
blanchet@41483
   259
            | NONE =>
blanchet@43314
   260
              0 |> fold (Integer.max
blanchet@43314
   261
                         o default_max_relevant_for_prover ctxt slicing)
blanchet@41483
   262
                        provers
blanchet@43862
   263
                |> mode = Auto_Try
blanchet@43862
   264
                   ? (fn n => n div auto_try_max_relevant_divisor)
blanchet@41483
   265
          val is_built_in_const =
blanchet@41483
   266
            is_built_in_const_for_prover ctxt (hd provers)
blanchet@41483
   267
        in
blanchet@43793
   268
          relevant_facts ctxt relevance_thresholds max_max_relevant
blanchet@43793
   269
                         is_appropriate_prop is_built_in_const relevance_fudge
blanchet@43793
   270
                         relevance_override chained_ths hyp_ts concl_t
blanchet@41483
   271
          |> tap (fn facts =>
blanchet@41483
   272
                     if debug then
blanchet@41483
   273
                       label ^ plural_s (length provers) ^ ": " ^
blanchet@41483
   274
                       (if null facts then
blanchet@41483
   275
                          "Found no relevant facts."
blanchet@41483
   276
                        else
blanchet@41483
   277
                          "Including (up to) " ^ string_of_int (length facts) ^
blanchet@41483
   278
                          " relevant fact" ^ plural_s (length facts) ^ ":\n" ^
blanchet@41483
   279
                          (facts |> map (fst o fst) |> space_implode " ") ^ ".")
blanchet@42644
   280
                       |> print
blanchet@41483
   281
                     else
blanchet@41483
   282
                       ())
blanchet@41483
   283
        end
blanchet@43793
   284
      fun launch_atps label is_appropriate_prop atps accum =
blanchet@43787
   285
        if null atps then
blanchet@41502
   286
          accum
blanchet@43793
   287
        else if not (is_appropriate_prop concl_t) then
blanchet@43787
   288
          (if verbose orelse length atps = length provers then
blanchet@43787
   289
             "Goal outside the scope of " ^
blanchet@43787
   290
             space_implode " " (serial_commas "and" (map quote atps)) ^ "."
blanchet@43787
   291
             |> Output.urgent_message
blanchet@43787
   292
           else
blanchet@43787
   293
             ();
blanchet@43787
   294
           accum)
blanchet@41502
   295
        else
blanchet@43785
   296
          launch_provers state
blanchet@43793
   297
              (get_facts label is_appropriate_prop atp_relevance_fudge o K atps)
blanchet@43785
   298
              (K (Untranslated_Fact o fst)) (K (K NONE)) atps
blanchet@42617
   299
      fun launch_smts accum =
blanchet@42617
   300
        if null smts then
blanchet@41483
   301
          accum
blanchet@41483
   302
        else
blanchet@41483
   303
          let
blanchet@43785
   304
            val facts = get_facts "SMT solver" (K true) smt_relevance_fudge smts
blanchet@43517
   305
            val weight = SMT_Weighted_Fact oo weight_smt_fact ctxt
blanchet@42612
   306
            fun smt_filter facts =
blanchet@42659
   307
              (if debug then curry (op o) SOME
blanchet@42659
   308
               else TimeLimit.timeLimit timeout o try)
boehmes@41680
   309
                  (SMT_Solver.smt_filter_preprocess state (facts ()))
blanchet@41483
   310
          in
blanchet@41483
   311
            smts |> map (`(class_of_smt_solver ctxt))
blanchet@41483
   312
                 |> AList.group (op =)
blanchet@43861
   313
                 |> map (snd #> launch_provers state (K facts) weight smt_filter
blanchet@43861
   314
                             #> fst)
blanchet@43861
   315
                 |> max_outcome_code |> rpair state
blanchet@41483
   316
          end
blanchet@43785
   317
      val launch_full_atps = launch_atps "ATP" (K true) full_atps
blanchet@43785
   318
      val launch_ueq_atps =
blanchet@43785
   319
        launch_atps "Unit equational provers" is_unit_equality ueq_atps
blanchet@41510
   320
      fun launch_atps_and_smt_solvers () =
blanchet@43785
   321
        [launch_full_atps, launch_ueq_atps, launch_smts]
blanchet@43862
   322
        |> smart_par_list_map (fn f => ignore (f (unknownN, state)))
blanchet@42644
   323
        handle ERROR msg => (print ("Error: " ^ msg); error msg)
blanchet@43862
   324
      fun maybe f (accum as (outcome_code, _)) =
blanchet@43862
   325
        accum |> (mode = Normal orelse outcome_code <> someN) ? f
blanchet@40241
   326
    in
blanchet@43861
   327
      (unknownN, state)
blanchet@43785
   328
      |> (if blocking then
blanchet@43862
   329
            launch_full_atps
blanchet@43862
   330
            #> mode <> Auto_Try ? (maybe launch_ueq_atps #> maybe launch_smts)
blanchet@43785
   331
          else
blanchet@43785
   332
            (fn p => Future.fork (tap launch_atps_and_smt_solvers) |> K p))
blanchet@42644
   333
      handle TimeLimit.TimeOut =>
blanchet@43861
   334
             (print "Sledgehammer ran out of time."; (unknownN, state))
blanchet@40241
   335
    end
blanchet@43861
   336
    |> `(fn (outcome_code, _) => outcome_code = someN)
blanchet@38290
   337
wenzelm@28582
   338
end;