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