src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Wed, 15 Dec 2010 11:26:29 +0100
changeset 41400 4a7410be4dfc
parent 41399 d58157bb1ae7
child 41407 1e12d6495423
permissions -rw-r--r--
crank up Metis's timeout for SMT solvers, since users love Metis
blanchet@41335
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_provers.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@41335
     6
Generic prover abstraction for Sledgehammer.
wenzelm@28477
     7
*)
wenzelm@28477
     8
blanchet@41335
     9
signature SLEDGEHAMMER_PROVERS =
wenzelm@28477
    10
sig
blanchet@40422
    11
  type failure = ATP_Proof.failure
blanchet@39232
    12
  type locality = Sledgehammer_Filter.locality
blanchet@40251
    13
  type relevance_fudge = Sledgehammer_Filter.relevance_fudge
blanchet@40358
    14
  type translated_formula = Sledgehammer_ATP_Translate.translated_formula
blanchet@41386
    15
  type type_system = Sledgehammer_ATP_Translate.type_system
blanchet@40249
    16
  type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
blanchet@39733
    17
blanchet@35967
    18
  type params =
blanchet@39226
    19
    {blocking: bool,
blanchet@39226
    20
     debug: bool,
blanchet@35967
    21
     verbose: bool,
blanchet@36141
    22
     overlord: bool,
blanchet@40240
    23
     provers: string list,
blanchet@41386
    24
     type_sys: type_system,
blanchet@36235
    25
     explicit_apply: bool,
blanchet@38984
    26
     relevance_thresholds: real * real,
blanchet@38983
    27
     max_relevant: int option,
blanchet@35967
    28
     isar_proof: bool,
blanchet@36916
    29
     isar_shrink_factor: int,
blanchet@39229
    30
     timeout: Time.time,
blanchet@39229
    31
     expect: string}
blanchet@39733
    32
blanchet@41338
    33
  datatype prover_fact =
blanchet@41338
    34
    Untranslated_Fact of (string * locality) * thm |
blanchet@41338
    35
    ATP_Translated_Fact of
blanchet@41339
    36
      translated_formula option * ((string * locality) * thm)
blanchet@40242
    37
blanchet@40242
    38
  type prover_problem =
blanchet@39564
    39
    {state: Proof.state,
blanchet@39242
    40
     goal: thm,
blanchet@39242
    41
     subgoal: int,
blanchet@40246
    42
     subgoal_count: int,
blanchet@41338
    43
     facts: prover_fact list}
blanchet@39733
    44
blanchet@40242
    45
  type prover_result =
blanchet@36370
    46
    {outcome: failure option,
blanchet@40445
    47
     used_facts: (string * locality) list,
blanchet@40243
    48
     run_time_in_msecs: int option,
blanchet@40243
    49
     message: string}
blanchet@39733
    50
blanchet@40242
    51
  type prover = params -> minimize_command -> prover_problem -> prover_result
blanchet@35867
    52
blanchet@41337
    53
  val das_Tool : string
blanchet@41335
    54
  val is_smt_prover : Proof.context -> string -> bool
blanchet@41189
    55
  val is_prover_available : Proof.context -> string -> bool
blanchet@40253
    56
  val is_prover_installed : Proof.context -> string -> bool
blanchet@41189
    57
  val default_max_relevant_for_prover : Proof.context -> string -> int
blanchet@40615
    58
  val is_built_in_const_for_prover :
blanchet@41314
    59
    Proof.context -> string -> string * typ -> term list -> bool
blanchet@41335
    60
  val atp_relevance_fudge : relevance_fudge
blanchet@41335
    61
  val smt_relevance_fudge : relevance_fudge
blanchet@41189
    62
  val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
blanchet@38257
    63
  val dest_dir : string Config.T
blanchet@38257
    64
  val problem_prefix : string Config.T
blanchet@39247
    65
  val measure_run_time : bool Config.T
blanchet@41339
    66
  val untranslated_fact : prover_fact -> (string * locality) * thm
blanchet@41189
    67
  val available_provers : Proof.context -> unit
blanchet@40240
    68
  val kill_provers : unit -> unit
blanchet@40240
    69
  val running_provers : unit -> unit
blanchet@40240
    70
  val messages : int option -> unit
blanchet@41189
    71
  val get_prover : Proof.context -> bool -> string -> prover
blanchet@38257
    72
  val setup : theory -> theory
wenzelm@28477
    73
end;
wenzelm@28477
    74
blanchet@41335
    75
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
wenzelm@28477
    76
struct
wenzelm@28477
    77
blanchet@38262
    78
open ATP_Problem
blanchet@39731
    79
open ATP_Proof
blanchet@38262
    80
open ATP_Systems
blanchet@39734
    81
open Metis_Translate
blanchet@38257
    82
open Sledgehammer_Util
blanchet@39232
    83
open Sledgehammer_Filter
blanchet@40249
    84
open Sledgehammer_ATP_Translate
blanchet@40249
    85
open Sledgehammer_ATP_Reconstruct
blanchet@37583
    86
blanchet@37583
    87
(** The Sledgehammer **)
blanchet@37583
    88
blanchet@38348
    89
(* Identifier to distinguish Sledgehammer from other tools using
blanchet@38348
    90
   "Async_Manager". *)
blanchet@37585
    91
val das_Tool = "Sledgehammer"
blanchet@37585
    92
blanchet@41189
    93
fun is_smt_prover ctxt name =
blanchet@41189
    94
  let val smts = SMT_Solver.available_solvers_of ctxt in
blanchet@41189
    95
    case try (unprefix remote_prefix) name of
blanchet@41189
    96
      SOME suffix => member (op =) smts suffix andalso
blanchet@41189
    97
                     SMT_Solver.is_remotely_available ctxt suffix
blanchet@41189
    98
    | NONE => member (op =) smts name
blanchet@41189
    99
  end
blanchet@40243
   100
blanchet@41189
   101
fun is_prover_available ctxt name =
blanchet@41189
   102
  let val thy = ProofContext.theory_of ctxt in
blanchet@41189
   103
    is_smt_prover ctxt name orelse member (op =) (available_atps thy) name
blanchet@41189
   104
  end
blanchet@40253
   105
blanchet@40253
   106
fun is_prover_installed ctxt name =
blanchet@40253
   107
  let val thy = ProofContext.theory_of ctxt in
blanchet@41189
   108
    if is_smt_prover ctxt name then
blanchet@41189
   109
      String.isPrefix remote_prefix name orelse
blanchet@41189
   110
      SMT_Solver.is_locally_installed ctxt name
blanchet@41189
   111
    else
blanchet@41189
   112
      is_atp_installed thy name
blanchet@41189
   113
  end
blanchet@41189
   114
blanchet@41189
   115
fun available_smt_solvers ctxt =
blanchet@41189
   116
  let val smts = SMT_Solver.available_solvers_of ctxt in
blanchet@41189
   117
    smts @ map (prefix remote_prefix)
blanchet@41189
   118
               (filter (SMT_Solver.is_remotely_available ctxt) smts)
blanchet@40253
   119
  end
blanchet@40253
   120
blanchet@41189
   121
fun default_max_relevant_for_prover ctxt name =
blanchet@41189
   122
  let val thy = ProofContext.theory_of ctxt in
blanchet@41228
   123
    if is_smt_prover ctxt name then
blanchet@41228
   124
      SMT_Solver.default_max_relevant ctxt
blanchet@41228
   125
          (perhaps (try (unprefix remote_prefix)) name)
blanchet@41228
   126
    else
blanchet@41228
   127
      #default_max_relevant (get_atp thy name)
blanchet@41189
   128
  end
blanchet@40244
   129
blanchet@41388
   130
(* These are either simplified away by "Meson.presimplify" (most of the time) or
blanchet@41388
   131
   handled specially via "fFalse", "fTrue", ..., "fequal". *)
blanchet@40252
   132
val atp_irrelevant_consts =
blanchet@41388
   133
  [@{const_name False}, @{const_name True}, @{const_name Not},
blanchet@41388
   134
   @{const_name conj}, @{const_name disj}, @{const_name implies},
blanchet@41388
   135
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
blanchet@40447
   136
blanchet@41314
   137
fun is_built_in_const_for_prover ctxt name (s, T) args =
blanchet@41314
   138
  if is_smt_prover ctxt name then SMT_Builtin.is_builtin_ext ctxt (s, T) args
blanchet@40615
   139
  else member (op =) atp_irrelevant_consts s
blanchet@40252
   140
blanchet@40251
   141
(* FUDGE *)
blanchet@40251
   142
val atp_relevance_fudge =
blanchet@40251
   143
  {worse_irrel_freq = 100.0,
blanchet@40251
   144
   higher_order_irrel_weight = 1.05,
blanchet@40251
   145
   abs_rel_weight = 0.5,
blanchet@40251
   146
   abs_irrel_weight = 2.0,
blanchet@40251
   147
   skolem_irrel_weight = 0.75,
blanchet@40251
   148
   theory_const_rel_weight = 0.5,
blanchet@40251
   149
   theory_const_irrel_weight = 0.25,
blanchet@40251
   150
   intro_bonus = 0.15,
blanchet@40251
   151
   elim_bonus = 0.15,
blanchet@40251
   152
   simp_bonus = 0.15,
blanchet@40251
   153
   local_bonus = 0.55,
blanchet@40251
   154
   assum_bonus = 1.05,
blanchet@40251
   155
   chained_bonus = 1.5,
blanchet@40251
   156
   max_imperfect = 11.5,
blanchet@40251
   157
   max_imperfect_exp = 1.0,
blanchet@40251
   158
   threshold_divisor = 2.0,
blanchet@41341
   159
   ridiculous_threshold = 0.01}
blanchet@40251
   160
blanchet@40252
   161
(* FUDGE (FIXME) *)
blanchet@40251
   162
val smt_relevance_fudge =
blanchet@40252
   163
  {worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
blanchet@40252
   164
   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
blanchet@40252
   165
   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
blanchet@40252
   166
   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
blanchet@40252
   167
   skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
blanchet@40252
   168
   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
blanchet@40252
   169
   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
blanchet@40252
   170
   intro_bonus = #intro_bonus atp_relevance_fudge,
blanchet@40252
   171
   elim_bonus = #elim_bonus atp_relevance_fudge,
blanchet@40252
   172
   simp_bonus = #simp_bonus atp_relevance_fudge,
blanchet@40252
   173
   local_bonus = #local_bonus atp_relevance_fudge,
blanchet@40252
   174
   assum_bonus = #assum_bonus atp_relevance_fudge,
blanchet@40252
   175
   chained_bonus = #chained_bonus atp_relevance_fudge,
blanchet@40252
   176
   max_imperfect = #max_imperfect atp_relevance_fudge,
blanchet@40252
   177
   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
blanchet@40252
   178
   threshold_divisor = #threshold_divisor atp_relevance_fudge,
blanchet@40252
   179
   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
blanchet@40251
   180
blanchet@41189
   181
fun relevance_fudge_for_prover ctxt name =
blanchet@41189
   182
  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
blanchet@40251
   183
blanchet@41189
   184
fun available_provers ctxt =
blanchet@40241
   185
  let
blanchet@41189
   186
    val thy = ProofContext.theory_of ctxt
blanchet@40241
   187
    val (remote_provers, local_provers) =
blanchet@41189
   188
      sort_strings (available_atps thy) @
blanchet@41189
   189
      sort_strings (available_smt_solvers ctxt)
blanchet@40241
   190
      |> List.partition (String.isPrefix remote_prefix)
blanchet@40241
   191
  in
blanchet@40446
   192
    Output.urgent_message ("Available provers: " ^
blanchet@40446
   193
                           commas (local_provers @ remote_provers) ^ ".")
blanchet@40241
   194
  end
blanchet@35967
   195
blanchet@40240
   196
fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
blanchet@40240
   197
fun running_provers () = Async_Manager.running_threads das_Tool "provers"
blanchet@40240
   198
val messages = Async_Manager.thread_messages das_Tool "prover"
blanchet@40240
   199
blanchet@40240
   200
(** problems, results, ATPs, etc. **)
blanchet@35967
   201
blanchet@35967
   202
type params =
blanchet@39226
   203
  {blocking: bool,
blanchet@39226
   204
   debug: bool,
blanchet@35967
   205
   verbose: bool,
blanchet@36141
   206
   overlord: bool,
blanchet@40240
   207
   provers: string list,
blanchet@41386
   208
   type_sys: type_system,
blanchet@36235
   209
   explicit_apply: bool,
blanchet@38984
   210
   relevance_thresholds: real * real,
blanchet@38983
   211
   max_relevant: int option,
blanchet@35967
   212
   isar_proof: bool,
blanchet@36916
   213
   isar_shrink_factor: int,
blanchet@39229
   214
   timeout: Time.time,
blanchet@39229
   215
   expect: string}
blanchet@35867
   216
blanchet@41338
   217
datatype prover_fact =
blanchet@41338
   218
  Untranslated_Fact of (string * locality) * thm |
blanchet@41339
   219
  ATP_Translated_Fact of translated_formula option * ((string * locality) * thm)
blanchet@40242
   220
blanchet@40242
   221
type prover_problem =
blanchet@39564
   222
  {state: Proof.state,
blanchet@39242
   223
   goal: thm,
blanchet@39242
   224
   subgoal: int,
blanchet@40246
   225
   subgoal_count: int,
blanchet@41338
   226
   facts: prover_fact list}
blanchet@35867
   227
blanchet@40242
   228
type prover_result =
blanchet@36370
   229
  {outcome: failure option,
blanchet@35967
   230
   message: string,
blanchet@40445
   231
   used_facts: (string * locality) list,
blanchet@40243
   232
   run_time_in_msecs: int option}
blanchet@35867
   233
blanchet@40242
   234
type prover = params -> minimize_command -> prover_problem -> prover_result
blanchet@35867
   235
blanchet@38257
   236
(* configuration attributes *)
wenzelm@28484
   237
blanchet@39235
   238
val (dest_dir, dest_dir_setup) =
blanchet@39247
   239
  Attrib.config_string "sledgehammer_dest_dir" (K "")
blanchet@39235
   240
  (* Empty string means create files in Isabelle's temporary files directory. *)
wenzelm@28484
   241
blanchet@38257
   242
val (problem_prefix, problem_prefix_setup) =
blanchet@39247
   243
  Attrib.config_string "sledgehammer_problem_prefix" (K "prob")
wenzelm@28477
   244
blanchet@39247
   245
val (measure_run_time, measure_run_time_setup) =
blanchet@39247
   246
  Attrib.config_bool "sledgehammer_measure_run_time" (K false)
wenzelm@32995
   247
blanchet@38257
   248
fun with_path cleanup after f path =
blanchet@38257
   249
  Exn.capture f path
blanchet@38257
   250
  |> tap (fn _ => cleanup path)
blanchet@38257
   251
  |> Exn.release
blanchet@38257
   252
  |> tap (after path)
wenzelm@28484
   253
blanchet@40241
   254
fun proof_banner auto =
blanchet@40465
   255
  if auto then "Auto Sledgehammer found a proof" else "Try this command"
blanchet@40241
   256
blanchet@40240
   257
(* generic TPTP-based ATPs *)
blanchet@38257
   258
blanchet@41339
   259
fun untranslated_fact (Untranslated_Fact p) = p
blanchet@41339
   260
  | untranslated_fact (ATP_Translated_Fact (_, p)) = p
blanchet@41338
   261
fun atp_translated_fact ctxt (Untranslated_Fact p) = translate_atp_fact ctxt p
blanchet@41339
   262
  | atp_translated_fact _ (ATP_Translated_Fact q) = q
blanchet@40242
   263
blanchet@40636
   264
fun int_opt_add (SOME m) (SOME n) = SOME (m + n)
blanchet@40636
   265
  | int_opt_add _ _ = NONE
blanchet@40243
   266
blanchet@39732
   267
(* Important messages are important but not so important that users want to see
blanchet@39732
   268
   them each time. *)
blanchet@39733
   269
val important_message_keep_factor = 0.1
blanchet@39732
   270
blanchet@40244
   271
fun run_atp auto atp_name
krauss@41351
   272
        ({exec, required_execs, arguments, has_incomplete_mode, proof_delims,
blanchet@41386
   273
          known_failures, explicit_forall, use_conjecture_for_hypotheses, ...}
blanchet@41386
   274
          : atp_config)
blanchet@41386
   275
        ({debug, verbose, overlord, type_sys, explicit_apply, isar_proof,
blanchet@41386
   276
          isar_shrink_factor, timeout, ...} : params)
blanchet@41229
   277
        minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
blanchet@38257
   278
  let
blanchet@39564
   279
    val ctxt = Proof.context_of state
blanchet@39242
   280
    val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
blanchet@41339
   281
    val facts = facts |> map (atp_translated_fact ctxt)
blanchet@40240
   282
    val dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
blanchet@40240
   283
                   else Config.get ctxt dest_dir
blanchet@40240
   284
    val problem_prefix = Config.get ctxt problem_prefix
blanchet@40242
   285
    val problem_file_name =
blanchet@39247
   286
      Path.basic ((if overlord then "prob_" ^ atp_name
blanchet@40240
   287
                   else problem_prefix ^ serial_string ())
blanchet@39247
   288
                  ^ "_" ^ string_of_int subgoal)
blanchet@40242
   289
    val problem_path_name =
blanchet@40240
   290
      if dest_dir = "" then
blanchet@40242
   291
        File.tmp_path problem_file_name
blanchet@40240
   292
      else if File.exists (Path.explode dest_dir) then
blanchet@40242
   293
        Path.append (Path.explode dest_dir) problem_file_name
blanchet@39247
   294
      else
blanchet@40240
   295
        error ("No such directory: " ^ quote dest_dir ^ ".")
blanchet@39247
   296
    val measure_run_time = verbose orelse Config.get ctxt measure_run_time
blanchet@38338
   297
    val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
blanchet@40240
   298
    (* write out problem file and call ATP *)
blanchet@38883
   299
    fun command_line complete timeout probfile =
blanchet@38257
   300
      let
blanchet@38257
   301
        val core = File.shell_path command ^ " " ^ arguments complete timeout ^
blanchet@38257
   302
                   " " ^ File.shell_path probfile
blanchet@38257
   303
      in
blanchet@39254
   304
        (if measure_run_time then "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }"
blanchet@38983
   305
         else "exec " ^ core) ^ " 2>&1"
blanchet@38257
   306
      end
blanchet@38257
   307
    fun split_time s =
blanchet@38257
   308
      let
blanchet@38257
   309
        val split = String.tokens (fn c => str c = "\n");
blanchet@38257
   310
        val (output, t) = s |> split |> split_last |> apfst cat_lines;
blanchet@38257
   311
        fun as_num f = f >> (fst o read_int);
blanchet@38257
   312
        val num = as_num (Scan.many1 Symbol.is_ascii_digit);
blanchet@38257
   313
        val digit = Scan.one Symbol.is_ascii_digit;
blanchet@38257
   314
        val num3 = as_num (digit ::: digit ::: (digit >> single));
blanchet@38257
   315
        val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b);
wenzelm@40875
   316
        val as_time = Scan.read Symbol.stopper time o raw_explode
blanchet@38257
   317
      in (output, as_time t) end;
blanchet@38257
   318
    fun run_on probfile =
blanchet@38338
   319
      case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
blanchet@38278
   320
        (home_var, _) :: _ =>
blanchet@38257
   321
        error ("The environment variable " ^ quote home_var ^ " is not set.")
blanchet@38278
   322
      | [] =>
blanchet@38278
   323
        if File.exists command then
blanchet@38278
   324
          let
blanchet@39564
   325
            fun run complete timeout =
blanchet@38278
   326
              let
blanchet@38883
   327
                val command = command_line complete timeout probfile
blanchet@38278
   328
                val ((output, msecs), res_code) =
blanchet@38278
   329
                  bash_output command
blanchet@38278
   330
                  |>> (if overlord then
blanchet@38278
   331
                         prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
blanchet@38278
   332
                       else
blanchet@38278
   333
                         I)
blanchet@40243
   334
                  |>> (if measure_run_time then split_time else rpair NONE)
blanchet@39692
   335
                val (tstplike_proof, outcome) =
blanchet@39692
   336
                  extract_tstplike_proof_and_outcome complete res_code
blanchet@39692
   337
                      proof_delims known_failures output
blanchet@39692
   338
              in (output, msecs, tstplike_proof, outcome) end
blanchet@38278
   339
            val readable_names = debug andalso overlord
blanchet@40445
   340
            val (atp_problem, pool, conjecture_offset, fact_names) =
blanchet@41382
   341
              prepare_atp_problem ctxt readable_names explicit_forall type_sys
blanchet@40445
   342
                                  explicit_apply hyp_ts concl_t facts
blanchet@39692
   343
            val ss = tptp_strings_for_atp_problem use_conjecture_for_hypotheses
blanchet@40240
   344
                                                  atp_problem
blanchet@38854
   345
            val _ = File.write_list probfile ss
blanchet@38278
   346
            val conjecture_shape =
blanchet@38278
   347
              conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
blanchet@38286
   348
              |> map single
blanchet@39564
   349
            val run_twice = has_incomplete_mode andalso not auto
blanchet@38883
   350
            val timer = Timer.startRealTimer ()
blanchet@38278
   351
            val result =
blanchet@39564
   352
              run false (if run_twice then
blanchet@39564
   353
                           Time.fromMilliseconds
blanchet@38883
   354
                                         (2 * Time.toMilliseconds timeout div 3)
blanchet@39564
   355
                         else
blanchet@39564
   356
                           timeout)
blanchet@39564
   357
              |> run_twice
blanchet@38883
   358
                 ? (fn (_, msecs0, _, SOME _) =>
blanchet@39564
   359
                       run true (Time.- (timeout, Timer.checkRealTimer timer))
blanchet@39692
   360
                       |> (fn (output, msecs, tstplike_proof, outcome) =>
blanchet@40636
   361
                              (output, int_opt_add msecs0 msecs, tstplike_proof,
blanchet@40636
   362
                               outcome))
blanchet@38883
   363
                     | result => result)
blanchet@40445
   364
          in ((pool, conjecture_shape, fact_names), result) end
blanchet@38278
   365
        else
blanchet@38278
   366
          error ("Bad executable: " ^ Path.implode command ^ ".")
blanchet@38257
   367
blanchet@38257
   368
    (* If the problem file has not been exported, remove it; otherwise, export
blanchet@38257
   369
       the proof file too. *)
blanchet@38257
   370
    fun cleanup probfile =
blanchet@40240
   371
      if dest_dir = "" then try File.rm probfile else NONE
blanchet@38257
   372
    fun export probfile (_, (output, _, _, _)) =
blanchet@40240
   373
      if dest_dir = "" then
blanchet@38257
   374
        ()
blanchet@38257
   375
      else
blanchet@38257
   376
        File.write (Path.explode (Path.implode probfile ^ "_proof")) output
blanchet@40445
   377
    val ((pool, conjecture_shape, fact_names),
blanchet@39692
   378
         (output, msecs, tstplike_proof, outcome)) =
blanchet@40242
   379
      with_path cleanup export run_on problem_path_name
blanchet@40445
   380
    val (conjecture_shape, fact_names) =
blanchet@40445
   381
      repair_conjecture_shape_and_fact_names output conjecture_shape fact_names
blanchet@39732
   382
    val important_message =
blanchet@40465
   383
      if not auto andalso random () <= important_message_keep_factor then
blanchet@39732
   384
        extract_important_message output
blanchet@39732
   385
      else
blanchet@39732
   386
        ""
blanchet@40445
   387
    val (message, used_facts) =
blanchet@38257
   388
      case outcome of
blanchet@38257
   389
        NONE =>
blanchet@38257
   390
        proof_text isar_proof
blanchet@38257
   391
            (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
blanchet@41384
   392
            (proof_banner auto, type_sys, minimize_command, tstplike_proof,
blanchet@40445
   393
             fact_names, goal, subgoal)
blanchet@38983
   394
        |>> (fn message =>
blanchet@40582
   395
                message ^
blanchet@40582
   396
                (if verbose then
blanchet@40582
   397
                   "\nATP real CPU time: " ^
blanchet@40582
   398
                   string_from_time (Time.fromMilliseconds (the msecs)) ^ "."
blanchet@40582
   399
                 else
blanchet@40582
   400
                   "") ^
blanchet@39353
   401
                (if important_message <> "" then
blanchet@39353
   402
                   "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^
blanchet@39353
   403
                   important_message
blanchet@39353
   404
                 else
blanchet@39353
   405
                   ""))
blanchet@40917
   406
      | SOME failure => (string_for_failure "ATP" failure, [])
blanchet@38257
   407
  in
blanchet@40445
   408
    {outcome = outcome, message = message, used_facts = used_facts,
blanchet@40242
   409
     run_time_in_msecs = msecs}
blanchet@38257
   410
  end
blanchet@38257
   411
blanchet@40917
   412
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
blanchet@40917
   413
   these are sorted out properly in the SMT module, we have to interpret these
blanchet@40917
   414
   ourselves. *)
blanchet@40932
   415
val remote_smt_failures =
blanchet@40932
   416
  [(22, CantConnect),
blanchet@40932
   417
   (127, NoPerl),
blanchet@40932
   418
   (2, NoLibwwwPerl)]
blanchet@40932
   419
val z3_failures =
blanchet@40932
   420
  [(103, MalformedInput),
blanchet@40932
   421
   (110, MalformedInput)]
blanchet@40932
   422
val unix_failures =
blanchet@40932
   423
  [(139, Crashed)]
blanchet@40932
   424
val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
blanchet@40797
   425
boehmes@40651
   426
fun failure_from_smt_failure (SMT_Failure.Counterexample _) = Unprovable
boehmes@40651
   427
  | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
boehmes@40802
   428
  | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
blanchet@40932
   429
    (case AList.lookup (op =) smt_failures code of
blanchet@40932
   430
       SOME failure => failure
blanchet@40932
   431
     | NONE => UnknownError)
boehmes@40651
   432
  | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
blanchet@40448
   433
  | failure_from_smt_failure _ = UnknownError
wenzelm@28586
   434
blanchet@40946
   435
(* FUDGE *)
blanchet@40795
   436
val smt_max_iter = 8
blanchet@40646
   437
val smt_iter_fact_divisor = 2
blanchet@40646
   438
val smt_iter_min_msecs = 5000
blanchet@40646
   439
val smt_iter_timeout_divisor = 2
blanchet@40666
   440
val smt_monomorph_limit = 4
blanchet@40636
   441
blanchet@40966
   442
fun smt_filter_loop ({verbose, timeout, ...} : params) remote state i =
blanchet@40636
   443
  let
blanchet@40795
   444
    val ctxt = Proof.context_of state
blanchet@40795
   445
    fun iter timeout iter_num outcome0 msecs_so_far facts =
blanchet@40795
   446
      let
blanchet@40795
   447
        val timer = Timer.startRealTimer ()
blanchet@40795
   448
        val ms = timeout |> Time.toMilliseconds
blanchet@40795
   449
        val iter_timeout =
blanchet@40795
   450
          if iter_num < smt_max_iter then
blanchet@40795
   451
            Int.min (ms, Int.max (smt_iter_min_msecs,
blanchet@40795
   452
                                  ms div smt_iter_timeout_divisor))
blanchet@40795
   453
            |> Time.fromMilliseconds
blanchet@40795
   454
          else
blanchet@40795
   455
            timeout
blanchet@40795
   456
        val num_facts = length facts
blanchet@40795
   457
        val _ =
blanchet@40795
   458
          if verbose then
blanchet@40795
   459
            "SMT iteration with " ^ string_of_int num_facts ^ " fact" ^
blanchet@40916
   460
            plural_s num_facts ^ " for " ^ string_from_time iter_timeout ^ "..."
blanchet@40795
   461
            |> Output.urgent_message
blanchet@40795
   462
          else
blanchet@40795
   463
            ()
blanchet@40795
   464
        val {outcome, used_facts, run_time_in_msecs} =
blanchet@41225
   465
          SMT_Solver.smt_filter remote iter_timeout state facts i
blanchet@40795
   466
        val _ =
blanchet@40795
   467
          if verbose andalso is_some outcome then
blanchet@40795
   468
            "SMT outcome: " ^ SMT_Failure.string_of_failure ctxt (the outcome)
blanchet@40795
   469
            |> Output.urgent_message
blanchet@40795
   470
          else
blanchet@40795
   471
            ()
blanchet@40795
   472
        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
blanchet@40795
   473
        val msecs_so_far = int_opt_add run_time_in_msecs msecs_so_far
blanchet@40795
   474
        val too_many_facts_perhaps =
blanchet@40795
   475
          case outcome of
blanchet@40795
   476
            NONE => false
blanchet@40795
   477
          | SOME (SMT_Failure.Counterexample _) => false
blanchet@40795
   478
          | SOME SMT_Failure.Time_Out => iter_timeout <> timeout
boehmes@40802
   479
          | SOME (SMT_Failure.Abnormal_Termination code) =>
blanchet@40795
   480
            (if verbose then
blanchet@40797
   481
               "The SMT solver invoked with " ^ string_of_int num_facts ^
blanchet@40940
   482
               " fact" ^ plural_s num_facts ^ " terminated abnormally with \
blanchet@40940
   483
               \exit code " ^ string_of_int code ^ "."
blanchet@40940
   484
               |> warning
blanchet@40795
   485
             else
blanchet@40795
   486
               ();
blanchet@40795
   487
             true (* kind of *))
blanchet@40795
   488
          | SOME SMT_Failure.Out_Of_Memory => true
blanchet@40795
   489
          | SOME _ => true
blanchet@40795
   490
        val timeout = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@40795
   491
      in
blanchet@40795
   492
        if too_many_facts_perhaps andalso iter_num < smt_max_iter andalso
blanchet@40795
   493
           num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
blanchet@40795
   494
          let val facts = take (num_facts div smt_iter_fact_divisor) facts in
blanchet@40795
   495
            iter timeout (iter_num + 1) outcome0 msecs_so_far facts
blanchet@40795
   496
          end
blanchet@40795
   497
        else
blanchet@40795
   498
          {outcome = if is_none outcome then NONE else the outcome0,
blanchet@40795
   499
           used_facts = used_facts, run_time_in_msecs = msecs_so_far}
blanchet@40636
   500
      end
blanchet@40795
   501
  in iter timeout 1 NONE (SOME 0) end
blanchet@40636
   502
blanchet@40914
   503
(* taken from "Mirabelle" and generalized *)
blanchet@40914
   504
fun can_apply timeout tac state i =
blanchet@40914
   505
  let
blanchet@40914
   506
    val {context = ctxt, facts, goal} = Proof.goal state
blanchet@40914
   507
    val full_tac = Method.insert_tac facts i THEN tac ctxt i
blanchet@40914
   508
  in
blanchet@40914
   509
    case try (TimeLimit.timeLimit timeout (Seq.pull o full_tac)) goal of
blanchet@40914
   510
      SOME (SOME _) => true
blanchet@40914
   511
    | _ => false
blanchet@40914
   512
  end
blanchet@40914
   513
blanchet@41400
   514
val smt_metis_timeout = seconds 1.0
blanchet@40914
   515
blanchet@40941
   516
fun can_apply_metis debug state i ths =
blanchet@40941
   517
  can_apply smt_metis_timeout
blanchet@40941
   518
            (Config.put Metis_Tactics.verbose debug
blanchet@40941
   519
             #> (fn ctxt => Metis_Tactics.metis_tac ctxt ths)) state i
blanchet@40914
   520
blanchet@41229
   521
fun run_smt_solver auto name (params as {debug, ...}) minimize_command
blanchet@41229
   522
        ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
blanchet@36379
   523
  let
blanchet@41189
   524
    val (remote, suffix) =
blanchet@41189
   525
      case try (unprefix remote_prefix) name of
blanchet@41189
   526
        SOME suffix => (true, suffix)
blanchet@41189
   527
      | NONE => (false, name)
blanchet@40666
   528
    val repair_context =
blanchet@41189
   529
      Context.proof_map (SMT_Config.select_solver suffix)
blanchet@41189
   530
      #> Config.put SMT_Config.verbose debug
blanchet@40666
   531
      #> Config.put SMT_Config.monomorph_limit smt_monomorph_limit
blanchet@40666
   532
    val state = state |> Proof.map_context repair_context
blanchet@40916
   533
    val thy = Proof.theory_of state
boehmes@41374
   534
    val facts = facts |> map (apsnd (pair NONE o Thm.transfer thy) o untranslated_fact)
blanchet@40422
   535
    val {outcome, used_facts, run_time_in_msecs} =
blanchet@41228
   536
      smt_filter_loop params remote state subgoal facts
blanchet@40966
   537
    val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
blanchet@40917
   538
    val outcome = outcome |> Option.map failure_from_smt_failure
blanchet@40244
   539
    val message =
blanchet@40425
   540
      case outcome of
blanchet@40425
   541
        NONE =>
blanchet@40914
   542
        let
blanchet@41399
   543
          val (method, settings) =
blanchet@40941
   544
            if can_apply_metis debug state subgoal (map snd used_facts) then
blanchet@41399
   545
              ("metis", "")
blanchet@40941
   546
            else
blanchet@41399
   547
              ("smt", if suffix = SMT_Config.solver_of @{context} then ""
blanchet@41399
   548
                      else "smt_solver = " ^ maybe_quote suffix)
blanchet@40914
   549
        in
blanchet@40914
   550
          try_command_line (proof_banner auto)
blanchet@41399
   551
              (apply_on_subgoal settings subgoal subgoal_count ^
blanchet@41399
   552
               command_call method (map fst other_lemmas)) ^
blanchet@40966
   553
          minimize_line minimize_command
blanchet@40966
   554
                        (map fst (other_lemmas @ chained_lemmas))
blanchet@40914
   555
        end
blanchet@40917
   556
      | SOME failure => string_for_failure "SMT solver" failure
blanchet@40244
   557
  in
blanchet@40914
   558
    {outcome = outcome, used_facts = map fst used_facts,
blanchet@40445
   559
     run_time_in_msecs = run_time_in_msecs, message = message}
blanchet@40244
   560
  end
blanchet@40244
   561
blanchet@41189
   562
fun get_prover ctxt auto name =
blanchet@41189
   563
  let val thy = ProofContext.theory_of ctxt in
blanchet@41189
   564
    if is_smt_prover ctxt name then
blanchet@41189
   565
      run_smt_solver auto name
blanchet@41189
   566
    else if member (op =) (available_atps thy) name then
blanchet@41189
   567
      run_atp auto name (get_atp thy name)
blanchet@41189
   568
    else
blanchet@41189
   569
      error ("No such prover: " ^ name ^ ".")
blanchet@41189
   570
  end
blanchet@40244
   571
blanchet@38257
   572
val setup =
blanchet@38257
   573
  dest_dir_setup
blanchet@38257
   574
  #> problem_prefix_setup
blanchet@39247
   575
  #> measure_run_time_setup
blanchet@38257
   576
wenzelm@28582
   577
end;