src/HOL/Tools/Sledgehammer/sledgehammer_run.ML
author blanchet
Wed, 08 Dec 2010 22:17:53 +0100
changeset 41339 0afdf5cde874
parent 41338 b98fe4de1ecd
child 41386 eb80538166b6
permissions -rw-r--r--
implicitly call the minimizer for SMT solvers that don't return an unsat core
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@41335
    13
  type params = Sledgehammer_Provers.params
blanchet@39733
    14
blanchet@38290
    15
  val run_sledgehammer :
blanchet@39564
    16
    params -> bool -> int -> relevance_override -> (string -> minimize_command)
blanchet@39564
    17
    -> Proof.state -> bool * Proof.state
wenzelm@28477
    18
end;
wenzelm@28477
    19
blanchet@41335
    20
structure Sledgehammer_Run : SLEDGEHAMMER_RUN =
wenzelm@28477
    21
struct
wenzelm@28477
    22
blanchet@38257
    23
open Sledgehammer_Util
blanchet@39232
    24
open Sledgehammer_Filter
blanchet@40249
    25
open Sledgehammer_ATP_Translate
blanchet@41335
    26
open Sledgehammer_Provers
blanchet@41339
    27
open Sledgehammer_Minimize
blanchet@40253
    28
blanchet@41337
    29
fun prover_description ctxt ({blocking, verbose, ...} : params) name num_facts i
blanchet@41337
    30
                       n goal =
blanchet@41337
    31
  quote name ^
blanchet@41337
    32
  (if verbose then
blanchet@41337
    33
     " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts
blanchet@41337
    34
   else
blanchet@41337
    35
     "") ^
blanchet@41337
    36
  " on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^ ":" ^
blanchet@41337
    37
  (if blocking then
blanchet@41337
    38
     ""
blanchet@41337
    39
   else
blanchet@41337
    40
     "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
blanchet@41337
    41
blanchet@41339
    42
val implicit_minimization_threshold = 50
blanchet@41339
    43
blanchet@41337
    44
fun run_prover (params as {blocking, debug, max_relevant, timeout, expect, ...})
blanchet@41337
    45
               auto minimize_command only
blanchet@41337
    46
               {state, goal, subgoal, subgoal_count, facts} name =
blanchet@41337
    47
  let
blanchet@41337
    48
    val ctxt = Proof.context_of state
blanchet@41337
    49
    val birth_time = Time.now ()
blanchet@41337
    50
    val death_time = Time.+ (birth_time, timeout)
blanchet@41337
    51
    val max_relevant =
blanchet@41337
    52
      the_default (default_max_relevant_for_prover ctxt name) max_relevant
blanchet@41337
    53
    val num_facts = length facts |> not only ? Integer.min max_relevant
blanchet@41337
    54
    val desc =
blanchet@41337
    55
      prover_description ctxt params name num_facts subgoal subgoal_count goal
blanchet@41337
    56
    val prover = get_prover ctxt auto name
blanchet@41337
    57
    val problem =
blanchet@41337
    58
      {state = state, goal = goal, subgoal = subgoal,
blanchet@41337
    59
       subgoal_count = subgoal_count, facts = take num_facts facts}
blanchet@41337
    60
    fun go () =
blanchet@41337
    61
      let
blanchet@41337
    62
        fun really_go () =
blanchet@41337
    63
          prover params (minimize_command name) problem
blanchet@41339
    64
          |> (fn {outcome, used_facts, message, ...} =>
blanchet@41339
    65
                 if is_some outcome then
blanchet@41339
    66
                   ("none", message)
blanchet@41339
    67
                 else
blanchet@41339
    68
                   ("some",
blanchet@41339
    69
                    if length used_facts >= implicit_minimization_threshold then
blanchet@41339
    70
                      minimize_facts params true subgoal subgoal_count state
blanchet@41339
    71
                          (filter_used_facts used_facts
blanchet@41339
    72
                               (map (apsnd single o untranslated_fact) facts))
blanchet@41339
    73
                      |> snd
blanchet@41339
    74
                    else
blanchet@41339
    75
                      message))
blanchet@41337
    76
        val (outcome_code, message) =
blanchet@41337
    77
          if debug then
blanchet@41337
    78
            really_go ()
blanchet@41337
    79
          else
blanchet@41337
    80
            (really_go ()
blanchet@41337
    81
             handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
blanchet@41337
    82
                  | exn =>
blanchet@41337
    83
                    if Exn.is_interrupt exn then
blanchet@41337
    84
                      reraise exn
blanchet@41337
    85
                    else
blanchet@41337
    86
                      ("unknown", "Internal error:\n" ^
blanchet@41337
    87
                                  ML_Compiler.exn_message exn ^ "\n"))
blanchet@41337
    88
        val _ =
blanchet@41337
    89
          if expect = "" orelse outcome_code = expect then
blanchet@41337
    90
            ()
blanchet@41337
    91
          else if blocking then
blanchet@41337
    92
            error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
blanchet@41337
    93
          else
blanchet@41337
    94
            warning ("Unexpected outcome: " ^ quote outcome_code ^ ".");
blanchet@41337
    95
      in (outcome_code = "some", message) end
blanchet@41337
    96
  in
blanchet@41337
    97
    if auto then
blanchet@41337
    98
      let val (success, message) = TimeLimit.timeLimit timeout go () in
blanchet@41337
    99
        (success, state |> success ? Proof.goal_message (fn () =>
blanchet@41339
   100
             Pretty.chunks [Pretty.str "",
blanchet@41339
   101
                            Pretty.mark Markup.hilite (Pretty.str message)]))
blanchet@41337
   102
      end
blanchet@41337
   103
    else if blocking then
blanchet@41337
   104
      let val (success, message) = TimeLimit.timeLimit timeout go () in
blanchet@41337
   105
        List.app Output.urgent_message
blanchet@41337
   106
                 (Async_Manager.break_into_chunks [desc ^ "\n" ^ message]);
blanchet@41337
   107
        (success, state)
blanchet@41337
   108
      end
blanchet@41337
   109
    else
blanchet@41337
   110
      (Async_Manager.launch das_Tool birth_time death_time desc (snd o go);
blanchet@41337
   111
       (false, state))
blanchet@41337
   112
  end
blanchet@41337
   113
blanchet@40946
   114
(* FUDGE *)
blanchet@40946
   115
val auto_max_relevant_divisor = 2
blanchet@40241
   116
blanchet@40446
   117
fun run_sledgehammer (params as {blocking, debug, provers, full_types,
blanchet@40250
   118
                                 relevance_thresholds, max_relevant, ...})
blanchet@39612
   119
                     auto i (relevance_override as {only, ...}) minimize_command
blanchet@39612
   120
                     state =
blanchet@40240
   121
  if null provers then
blanchet@40240
   122
    error "No prover is set."
blanchet@39564
   123
  else case subgoal_count state of
wenzelm@40392
   124
    0 => (Output.urgent_message "No subgoal!"; (false, state))
blanchet@39564
   125
  | n =>
blanchet@39564
   126
    let
blanchet@39610
   127
      val _ = Proof.assert_backward state
blanchet@40441
   128
      val ctxt = Proof.context_of state
blanchet@40441
   129
      val {facts = chained_ths, goal, ...} = Proof.goal state
blanchet@40241
   130
      val (_, hyp_ts, concl_t) = strip_subgoal goal i
blanchet@40240
   131
      val _ = () |> not blocking ? kill_provers
blanchet@41189
   132
      val _ = case find_first (not o is_prover_available ctxt) provers of
blanchet@41189
   133
                SOME name => error ("No such prover: " ^ name ^ ".")
blanchet@41189
   134
              | NONE => ()
wenzelm@40392
   135
      val _ = if auto then () else Output.urgent_message "Sledgehammering..."
blanchet@41189
   136
      val (smts, atps) = provers |> List.partition (is_smt_prover ctxt)
blanchet@40615
   137
      fun run_provers label full_types relevance_fudge maybe_translate provers
blanchet@40615
   138
                      (res as (success, state)) =
blanchet@40246
   139
        if success orelse null provers then
blanchet@40241
   140
          res
blanchet@40241
   141
        else
blanchet@40241
   142
          let
blanchet@40241
   143
            val max_max_relevant =
blanchet@40241
   144
              case max_relevant of
blanchet@40241
   145
                SOME n => n
blanchet@40241
   146
              | NONE =>
blanchet@41189
   147
                0 |> fold (Integer.max o default_max_relevant_for_prover ctxt)
blanchet@40246
   148
                          provers
blanchet@40241
   149
                  |> auto ? (fn n => n div auto_max_relevant_divisor)
blanchet@40615
   150
            val is_built_in_const =
blanchet@40615
   151
              is_built_in_const_for_prover ctxt (hd provers)
blanchet@40445
   152
            val facts =
blanchet@40246
   153
              relevant_facts ctxt full_types relevance_thresholds
blanchet@40615
   154
                             max_max_relevant is_built_in_const relevance_fudge
blanchet@40252
   155
                             relevance_override chained_ths hyp_ts concl_t
blanchet@40358
   156
              |> map maybe_translate
blanchet@40243
   157
            val problem =
blanchet@40246
   158
              {state = state, goal = goal, subgoal = i, subgoal_count = n,
blanchet@41229
   159
               facts = facts}
blanchet@41229
   160
            val run_prover = run_prover params auto minimize_command only
blanchet@40241
   161
          in
blanchet@40446
   162
            if debug then
blanchet@40616
   163
              Output.urgent_message (label ^ plural_s (length provers) ^ ": " ^
blanchet@40446
   164
                  (if null facts then
blanchet@40446
   165
                     "Found no relevant facts."
blanchet@40446
   166
                   else
blanchet@40446
   167
                     "Including (up to) " ^ string_of_int (length facts) ^
blanchet@40446
   168
                     " relevant fact" ^ plural_s (length facts) ^ ":\n" ^
blanchet@41339
   169
                     (facts |> map (fst o fst o untranslated_fact)
blanchet@40446
   170
                            |> space_implode " ") ^ "."))
blanchet@40446
   171
            else
blanchet@40446
   172
              ();
blanchet@40241
   173
            if auto then
blanchet@40242
   174
              fold (fn prover => fn (true, state) => (true, state)
blanchet@40245
   175
                                  | (false, _) => run_prover problem prover)
blanchet@40246
   176
                   provers (false, state)
blanchet@40241
   177
            else
blanchet@40246
   178
              provers |> (if blocking then Par_List.map else map)
blanchet@40246
   179
                             (run_prover problem)
blanchet@40246
   180
                      |> exists fst |> rpair state
blanchet@40241
   181
          end
blanchet@40252
   182
      val run_atps =
blanchet@40616
   183
        run_provers "ATP" full_types atp_relevance_fudge
blanchet@41338
   184
                    (ATP_Translated_Fact o translate_atp_fact ctxt) atps
blanchet@40252
   185
      val run_smts =
blanchet@41338
   186
        run_provers "SMT solver" true smt_relevance_fudge Untranslated_Fact smts
blanchet@40241
   187
      fun run_atps_and_smt_solvers () =
blanchet@40246
   188
        [run_atps, run_smts] |> Par_List.map (fn f => f (false, state) |> K ())
blanchet@40241
   189
    in
blanchet@40246
   190
      (false, state)
blanchet@40246
   191
      |> (if blocking then run_atps #> not auto ? run_smts
blanchet@40246
   192
          else (fn p => Future.fork (tap run_atps_and_smt_solvers) |> K p))
blanchet@40241
   193
    end
blanchet@38290
   194
wenzelm@28582
   195
end;