src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Fri, 06 May 2011 13:34:59 +0200
changeset 43580 e7af132d48fe
parent 43569 d4f5fec71ded
child 43587 626e292d22a7
permissions -rw-r--r--
allow each prover to specify its own formula kind for symbols occurring in the conjecture
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@43450
    18
  datatype rich_type_system =
blanchet@43450
    19
    Dumb_Type_Sys of type_system |
blanchet@43450
    20
    Smart_Type_Sys of bool
blanchet@43450
    21
blanchet@35967
    22
  type params =
blanchet@41456
    23
    {debug: bool,
blanchet@35967
    24
     verbose: bool,
blanchet@36141
    25
     overlord: bool,
blanchet@41456
    26
     blocking: bool,
blanchet@40240
    27
     provers: string list,
blanchet@43450
    28
     type_sys: rich_type_system,
blanchet@43051
    29
     relevance_thresholds: real * real,
blanchet@43051
    30
     max_relevant: int option,
blanchet@43051
    31
     monomorphize_limit: int,
blanchet@36235
    32
     explicit_apply: bool,
blanchet@35967
    33
     isar_proof: bool,
blanchet@36916
    34
     isar_shrink_factor: int,
blanchet@43314
    35
     slicing: bool,
blanchet@39229
    36
     timeout: Time.time,
blanchet@39229
    37
     expect: string}
blanchet@39733
    38
blanchet@41338
    39
  datatype prover_fact =
blanchet@41338
    40
    Untranslated_Fact of (string * locality) * thm |
blanchet@41483
    41
    SMT_Weighted_Fact of (string * locality) * (int option * thm)
blanchet@40242
    42
blanchet@40242
    43
  type prover_problem =
blanchet@39564
    44
    {state: Proof.state,
blanchet@39242
    45
     goal: thm,
blanchet@39242
    46
     subgoal: int,
blanchet@40246
    47
     subgoal_count: int,
blanchet@41483
    48
     facts: prover_fact list,
blanchet@42612
    49
     smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
blanchet@39733
    50
blanchet@40242
    51
  type prover_result =
blanchet@36370
    52
    {outcome: failure option,
blanchet@40445
    53
     used_facts: (string * locality) list,
blanchet@40243
    54
     run_time_in_msecs: int option,
blanchet@40243
    55
     message: string}
blanchet@39733
    56
blanchet@40242
    57
  type prover = params -> minimize_command -> prover_problem -> prover_result
blanchet@35867
    58
blanchet@43517
    59
  val smt_triggers : bool Config.T
blanchet@43517
    60
  val smt_weights : bool Config.T
blanchet@43517
    61
  val smt_weight_min_facts : int Config.T
blanchet@43517
    62
  val smt_min_weight : int Config.T
blanchet@43517
    63
  val smt_max_weight : int Config.T
blanchet@43517
    64
  val smt_max_weight_index : int Config.T
blanchet@41502
    65
  val smt_weight_curve : (int -> int) Unsynchronized.ref
blanchet@43517
    66
  val smt_max_slices : int Config.T
blanchet@43517
    67
  val smt_slice_fact_frac : real Config.T
blanchet@43517
    68
  val smt_slice_time_frac : real Config.T
blanchet@43517
    69
  val smt_slice_min_secs : int Config.T
blanchet@41337
    70
  val das_Tool : string
blanchet@41483
    71
  val select_smt_solver : string -> Proof.context -> Proof.context
blanchet@41335
    72
  val is_smt_prover : Proof.context -> string -> bool
blanchet@42591
    73
  val is_prover_supported : Proof.context -> string -> bool
blanchet@40253
    74
  val is_prover_installed : Proof.context -> string -> bool
blanchet@43314
    75
  val default_max_relevant_for_prover : Proof.context -> bool -> string -> int
blanchet@40615
    76
  val is_built_in_const_for_prover :
blanchet@41584
    77
    Proof.context -> string -> string * typ -> term list -> bool * term list
blanchet@41335
    78
  val atp_relevance_fudge : relevance_fudge
blanchet@41335
    79
  val smt_relevance_fudge : relevance_fudge
blanchet@41189
    80
  val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
blanchet@38257
    81
  val dest_dir : string Config.T
blanchet@38257
    82
  val problem_prefix : string Config.T
blanchet@39247
    83
  val measure_run_time : bool Config.T
blanchet@41502
    84
  val weight_smt_fact :
blanchet@43517
    85
    Proof.context -> int -> ((string * locality) * thm) * int
blanchet@41502
    86
    -> (string * locality) * (int option * thm)
blanchet@41339
    87
  val untranslated_fact : prover_fact -> (string * locality) * thm
blanchet@41483
    88
  val smt_weighted_fact :
blanchet@43517
    89
    Proof.context -> int -> prover_fact * int
blanchet@41502
    90
    -> (string * locality) * (int option * thm)
blanchet@42591
    91
  val supported_provers : Proof.context -> unit
blanchet@40240
    92
  val kill_provers : unit -> unit
blanchet@40240
    93
  val running_provers : unit -> unit
blanchet@40240
    94
  val messages : int option -> unit
blanchet@43315
    95
  val get_prover : Proof.context -> bool -> string -> prover
wenzelm@28477
    96
end;
wenzelm@28477
    97
blanchet@41335
    98
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
wenzelm@28477
    99
struct
wenzelm@28477
   100
blanchet@38262
   101
open ATP_Problem
blanchet@39731
   102
open ATP_Proof
blanchet@38262
   103
open ATP_Systems
blanchet@39734
   104
open Metis_Translate
blanchet@38257
   105
open Sledgehammer_Util
blanchet@39232
   106
open Sledgehammer_Filter
blanchet@40249
   107
open Sledgehammer_ATP_Translate
blanchet@40249
   108
open Sledgehammer_ATP_Reconstruct
blanchet@37583
   109
blanchet@37583
   110
(** The Sledgehammer **)
blanchet@37583
   111
blanchet@38348
   112
(* Identifier to distinguish Sledgehammer from other tools using
blanchet@38348
   113
   "Async_Manager". *)
blanchet@37585
   114
val das_Tool = "Sledgehammer"
blanchet@37585
   115
blanchet@41483
   116
val select_smt_solver =
boehmes@41680
   117
  Context.proof_map o SMT_Config.select_solver
blanchet@41483
   118
blanchet@41189
   119
fun is_smt_prover ctxt name =
boehmes@41680
   120
  member (op =) (SMT_Solver.available_solvers_of ctxt) name
blanchet@40243
   121
blanchet@42591
   122
fun is_prover_supported ctxt name =
wenzelm@43232
   123
  let val thy = Proof_Context.theory_of ctxt in
blanchet@42591
   124
    is_smt_prover ctxt name orelse member (op =) (supported_atps thy) name
blanchet@41189
   125
  end
blanchet@40253
   126
boehmes@41680
   127
fun is_prover_installed ctxt =
wenzelm@43232
   128
  is_smt_prover ctxt orf is_atp_installed (Proof_Context.theory_of ctxt)
blanchet@41189
   129
blanchet@43314
   130
fun get_slices slicing slices =
blanchet@43318
   131
  (0 upto length slices - 1) ~~ slices |> not slicing ? (List.last #> single)
blanchet@43314
   132
blanchet@43314
   133
fun default_max_relevant_for_prover ctxt slicing name =
wenzelm@43232
   134
  let val thy = Proof_Context.theory_of ctxt in
blanchet@41228
   135
    if is_smt_prover ctxt name then
boehmes@41680
   136
      SMT_Solver.default_max_relevant ctxt name
blanchet@41228
   137
    else
blanchet@43314
   138
      fold (Integer.max o snd o snd o snd)
blanchet@43517
   139
           (get_slices slicing (#best_slices (get_atp thy name) ctxt)) 0
blanchet@41189
   140
  end
blanchet@40244
   141
blanchet@41388
   142
(* These are either simplified away by "Meson.presimplify" (most of the time) or
blanchet@41388
   143
   handled specially via "fFalse", "fTrue", ..., "fequal". *)
blanchet@40252
   144
val atp_irrelevant_consts =
blanchet@41388
   145
  [@{const_name False}, @{const_name True}, @{const_name Not},
blanchet@41388
   146
   @{const_name conj}, @{const_name disj}, @{const_name implies},
blanchet@41388
   147
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
blanchet@40447
   148
blanchet@41483
   149
fun is_built_in_const_for_prover ctxt name =
blanchet@41483
   150
  if is_smt_prover ctxt name then
blanchet@41584
   151
    let val ctxt = ctxt |> select_smt_solver name in
blanchet@41584
   152
      fn x => fn ts =>
blanchet@41584
   153
         if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
blanchet@41584
   154
           (true, [])
blanchet@41584
   155
         else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
blanchet@41584
   156
           (true, ts)
blanchet@41584
   157
         else
blanchet@41584
   158
           (false, ts)
blanchet@41584
   159
    end
blanchet@41483
   160
  else
blanchet@41584
   161
    fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts)
blanchet@40252
   162
blanchet@40251
   163
(* FUDGE *)
blanchet@40251
   164
val atp_relevance_fudge =
blanchet@41407
   165
  {allow_ext = true,
blanchet@42661
   166
   local_const_multiplier = 1.5,
blanchet@41407
   167
   worse_irrel_freq = 100.0,
blanchet@40251
   168
   higher_order_irrel_weight = 1.05,
blanchet@40251
   169
   abs_rel_weight = 0.5,
blanchet@40251
   170
   abs_irrel_weight = 2.0,
blanchet@40251
   171
   skolem_irrel_weight = 0.75,
blanchet@40251
   172
   theory_const_rel_weight = 0.5,
blanchet@40251
   173
   theory_const_irrel_weight = 0.25,
blanchet@40251
   174
   intro_bonus = 0.15,
blanchet@40251
   175
   elim_bonus = 0.15,
blanchet@40251
   176
   simp_bonus = 0.15,
blanchet@40251
   177
   local_bonus = 0.55,
blanchet@40251
   178
   assum_bonus = 1.05,
blanchet@40251
   179
   chained_bonus = 1.5,
blanchet@40251
   180
   max_imperfect = 11.5,
blanchet@40251
   181
   max_imperfect_exp = 1.0,
blanchet@40251
   182
   threshold_divisor = 2.0,
blanchet@41341
   183
   ridiculous_threshold = 0.01}
blanchet@40251
   184
blanchet@40252
   185
(* FUDGE (FIXME) *)
blanchet@40251
   186
val smt_relevance_fudge =
blanchet@41407
   187
  {allow_ext = false,
blanchet@42661
   188
   local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
blanchet@41407
   189
   worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
blanchet@40252
   190
   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
blanchet@40252
   191
   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
blanchet@40252
   192
   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
blanchet@40252
   193
   skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
blanchet@40252
   194
   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
blanchet@40252
   195
   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
blanchet@40252
   196
   intro_bonus = #intro_bonus atp_relevance_fudge,
blanchet@40252
   197
   elim_bonus = #elim_bonus atp_relevance_fudge,
blanchet@40252
   198
   simp_bonus = #simp_bonus atp_relevance_fudge,
blanchet@40252
   199
   local_bonus = #local_bonus atp_relevance_fudge,
blanchet@40252
   200
   assum_bonus = #assum_bonus atp_relevance_fudge,
blanchet@40252
   201
   chained_bonus = #chained_bonus atp_relevance_fudge,
blanchet@40252
   202
   max_imperfect = #max_imperfect atp_relevance_fudge,
blanchet@40252
   203
   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
blanchet@40252
   204
   threshold_divisor = #threshold_divisor atp_relevance_fudge,
blanchet@40252
   205
   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
blanchet@40251
   206
blanchet@41189
   207
fun relevance_fudge_for_prover ctxt name =
blanchet@41189
   208
  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
blanchet@40251
   209
blanchet@42591
   210
fun supported_provers ctxt =
blanchet@40241
   211
  let
wenzelm@43232
   212
    val thy = Proof_Context.theory_of ctxt
blanchet@40241
   213
    val (remote_provers, local_provers) =
blanchet@42591
   214
      sort_strings (supported_atps thy) @
blanchet@42591
   215
      sort_strings (SMT_Solver.available_solvers_of ctxt)
blanchet@40241
   216
      |> List.partition (String.isPrefix remote_prefix)
blanchet@40241
   217
  in
blanchet@42591
   218
    Output.urgent_message ("Supported provers: " ^
blanchet@40446
   219
                           commas (local_provers @ remote_provers) ^ ".")
blanchet@40241
   220
  end
blanchet@35967
   221
blanchet@40240
   222
fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
blanchet@40240
   223
fun running_provers () = Async_Manager.running_threads das_Tool "provers"
blanchet@40240
   224
val messages = Async_Manager.thread_messages das_Tool "prover"
blanchet@40240
   225
blanchet@40240
   226
(** problems, results, ATPs, etc. **)
blanchet@35967
   227
blanchet@43450
   228
datatype rich_type_system =
blanchet@43450
   229
  Dumb_Type_Sys of type_system |
blanchet@43450
   230
  Smart_Type_Sys of bool
blanchet@43450
   231
blanchet@35967
   232
type params =
blanchet@41456
   233
  {debug: bool,
blanchet@35967
   234
   verbose: bool,
blanchet@36141
   235
   overlord: bool,
blanchet@41456
   236
   blocking: bool,
blanchet@40240
   237
   provers: string list,
blanchet@43450
   238
   type_sys: rich_type_system,
blanchet@43051
   239
   relevance_thresholds: real * real,
blanchet@43051
   240
   max_relevant: int option,
blanchet@43051
   241
   monomorphize_limit: int,
blanchet@36235
   242
   explicit_apply: bool,
blanchet@35967
   243
   isar_proof: bool,
blanchet@36916
   244
   isar_shrink_factor: int,
blanchet@43314
   245
   slicing: bool,
blanchet@39229
   246
   timeout: Time.time,
blanchet@39229
   247
   expect: string}
blanchet@35867
   248
blanchet@41338
   249
datatype prover_fact =
blanchet@41338
   250
  Untranslated_Fact of (string * locality) * thm |
blanchet@41483
   251
  SMT_Weighted_Fact of (string * locality) * (int option * thm)
blanchet@40242
   252
blanchet@40242
   253
type prover_problem =
blanchet@39564
   254
  {state: Proof.state,
blanchet@39242
   255
   goal: thm,
blanchet@39242
   256
   subgoal: int,
blanchet@40246
   257
   subgoal_count: int,
blanchet@41483
   258
   facts: prover_fact list,
blanchet@42612
   259
   smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
blanchet@35867
   260
blanchet@40242
   261
type prover_result =
blanchet@36370
   262
  {outcome: failure option,
blanchet@35967
   263
   message: string,
blanchet@40445
   264
   used_facts: (string * locality) list,
blanchet@40243
   265
   run_time_in_msecs: int option}
blanchet@35867
   266
blanchet@40242
   267
type prover = params -> minimize_command -> prover_problem -> prover_result
blanchet@35867
   268
blanchet@38257
   269
(* configuration attributes *)
wenzelm@28484
   270
wenzelm@43487
   271
val dest_dir =
wenzelm@43487
   272
  Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
blanchet@39235
   273
  (* Empty string means create files in Isabelle's temporary files directory. *)
wenzelm@28484
   274
wenzelm@43487
   275
val problem_prefix =
wenzelm@43487
   276
  Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
wenzelm@28477
   277
wenzelm@43487
   278
val measure_run_time =
wenzelm@43487
   279
  Attrib.setup_config_bool @{binding sledgehammer_measure_run_time} (K false)
wenzelm@32995
   280
blanchet@38257
   281
fun with_path cleanup after f path =
blanchet@38257
   282
  Exn.capture f path
blanchet@38257
   283
  |> tap (fn _ => cleanup path)
blanchet@38257
   284
  |> Exn.release
blanchet@38257
   285
  |> tap (after path)
wenzelm@28484
   286
blanchet@40241
   287
fun proof_banner auto =
blanchet@40465
   288
  if auto then "Auto Sledgehammer found a proof" else "Try this command"
blanchet@40241
   289
blanchet@43517
   290
val smt_triggers =
blanchet@43517
   291
  Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
blanchet@43517
   292
val smt_weights =
blanchet@43517
   293
  Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
blanchet@43517
   294
val smt_weight_min_facts =
blanchet@43517
   295
  Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
blanchet@41502
   296
blanchet@41502
   297
(* FUDGE *)
blanchet@43517
   298
val smt_min_weight =
blanchet@43517
   299
  Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
blanchet@43517
   300
val smt_max_weight =
blanchet@43517
   301
  Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
blanchet@43517
   302
val smt_max_weight_index =
blanchet@43517
   303
  Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
blanchet@41502
   304
val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
blanchet@41502
   305
blanchet@43517
   306
fun smt_fact_weight ctxt j num_facts =
blanchet@43517
   307
  if Config.get ctxt smt_weights andalso
blanchet@43517
   308
     num_facts >= Config.get ctxt smt_weight_min_facts then
blanchet@43517
   309
    let
blanchet@43517
   310
      val min = Config.get ctxt smt_min_weight
blanchet@43517
   311
      val max = Config.get ctxt smt_max_weight
blanchet@43517
   312
      val max_index = Config.get ctxt smt_max_weight_index
blanchet@43517
   313
      val curve = !smt_weight_curve
blanchet@43517
   314
    in
blanchet@43517
   315
      SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
blanchet@43517
   316
            div curve max_index)
blanchet@43517
   317
    end
blanchet@41502
   318
  else
blanchet@41502
   319
    NONE
blanchet@41502
   320
blanchet@43517
   321
fun weight_smt_fact ctxt num_facts ((info, th), j) =
blanchet@43517
   322
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43517
   323
    (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
blanchet@43517
   324
  end
blanchet@38257
   325
blanchet@41339
   326
fun untranslated_fact (Untranslated_Fact p) = p
blanchet@41483
   327
  | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
blanchet@43450
   328
fun atp_translated_fact ctxt fact =
blanchet@43415
   329
    translate_atp_fact ctxt false (untranslated_fact fact)
blanchet@41502
   330
fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
blanchet@43517
   331
  | smt_weighted_fact ctxt num_facts (fact, j) =
blanchet@43517
   332
    (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts
blanchet@41502
   333
blanchet@41561
   334
fun overlord_file_location_for_prover prover =
blanchet@41561
   335
  (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
blanchet@41561
   336
blanchet@41561
   337
blanchet@41502
   338
(* generic TPTP-based ATPs *)
blanchet@40242
   339
blanchet@40636
   340
fun int_opt_add (SOME m) (SOME n) = SOME (m + n)
blanchet@40636
   341
  | int_opt_add _ _ = NONE
blanchet@40243
   342
blanchet@43323
   343
val atp_blacklist_max_iters = 10
blanchet@39732
   344
(* Important messages are important but not so important that users want to see
blanchet@39732
   345
   them each time. *)
blanchet@43480
   346
val atp_important_message_keep_quotient = 10
blanchet@39732
   347
blanchet@43460
   348
val fallback_best_type_systems =
blanchet@43554
   349
  [Preds (Polymorphic, Const_Arg_Types), Simple All_Types]
blanchet@43450
   350
blanchet@43554
   351
fun adjust_dumb_type_sys formats (Simple level) =
blanchet@43554
   352
    if member (op =) formats Tff then (Tff, Simple level)
blanchet@43552
   353
    else (Fof, Preds (Mangled_Monomorphic, level))
blanchet@43450
   354
  | adjust_dumb_type_sys formats type_sys =
blanchet@43450
   355
    if member (op =) formats Fof then (Fof, type_sys)
blanchet@43554
   356
    else (Tff, Simple (level_of_type_sys type_sys))
blanchet@43450
   357
fun determine_format_and_type_sys _ formats (Dumb_Type_Sys type_sys) =
blanchet@43450
   358
    adjust_dumb_type_sys formats type_sys
blanchet@43450
   359
  | determine_format_and_type_sys best_type_systems formats
blanchet@43450
   360
                                  (Smart_Type_Sys full_types) =
blanchet@43484
   361
    map type_sys_from_string best_type_systems @ fallback_best_type_systems
blanchet@43460
   362
    |> find_first (if full_types then is_type_sys_virtually_sound else K true)
blanchet@43460
   363
    |> the |> adjust_dumb_type_sys formats
blanchet@43419
   364
blanchet@43315
   365
fun run_atp auto name
blanchet@43449
   366
        ({exec, required_execs, arguments, proof_delims, known_failures,
blanchet@43580
   367
          conj_sym_kind, prem_kind, formats, best_slices, best_type_systems,
blanchet@43580
   368
          ...} : atp_config)
blanchet@43450
   369
        ({debug, verbose, overlord, type_sys, max_relevant, monomorphize_limit,
blanchet@43450
   370
          explicit_apply, isar_proof, isar_shrink_factor, slicing, timeout, ...}
blanchet@43450
   371
         : params)
blanchet@41229
   372
        minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
blanchet@38257
   373
  let
blanchet@43053
   374
    val thy = Proof.theory_of state
blanchet@39564
   375
    val ctxt = Proof.context_of state
blanchet@43450
   376
    val (format, type_sys) =
blanchet@43450
   377
      determine_format_and_type_sys best_type_systems formats type_sys
blanchet@39242
   378
    val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
blanchet@41407
   379
    val (dest_dir, problem_prefix) =
blanchet@41407
   380
      if overlord then overlord_file_location_for_prover name
blanchet@41407
   381
      else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
blanchet@40242
   382
    val problem_file_name =
blanchet@41407
   383
      Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
blanchet@41407
   384
                  "_" ^ string_of_int subgoal)
blanchet@40242
   385
    val problem_path_name =
blanchet@40240
   386
      if dest_dir = "" then
blanchet@40242
   387
        File.tmp_path problem_file_name
blanchet@40240
   388
      else if File.exists (Path.explode dest_dir) then
blanchet@40242
   389
        Path.append (Path.explode dest_dir) problem_file_name
blanchet@39247
   390
      else
blanchet@40240
   391
        error ("No such directory: " ^ quote dest_dir ^ ".")
blanchet@39247
   392
    val measure_run_time = verbose orelse Config.get ctxt measure_run_time
blanchet@38338
   393
    val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
blanchet@38257
   394
    fun split_time s =
blanchet@38257
   395
      let
blanchet@43319
   396
        val split = String.tokens (fn c => str c = "\n")
blanchet@43319
   397
        val (output, t) = s |> split |> split_last |> apfst cat_lines
blanchet@43319
   398
        fun as_num f = f >> (fst o read_int)
blanchet@43319
   399
        val num = as_num (Scan.many1 Symbol.is_ascii_digit)
blanchet@43319
   400
        val digit = Scan.one Symbol.is_ascii_digit
blanchet@43319
   401
        val num3 = as_num (digit ::: digit ::: (digit >> single))
blanchet@43319
   402
        val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
wenzelm@40875
   403
        val as_time = Scan.read Symbol.stopper time o raw_explode
blanchet@43319
   404
      in (output, as_time t) end
blanchet@41561
   405
    fun run_on prob_file =
blanchet@38338
   406
      case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
blanchet@38278
   407
        (home_var, _) :: _ =>
blanchet@38257
   408
        error ("The environment variable " ^ quote home_var ^ " is not set.")
blanchet@38278
   409
      | [] =>
blanchet@38278
   410
        if File.exists command then
blanchet@38278
   411
          let
blanchet@43314
   412
            (* If slicing is disabled, we expand the last slice to fill the
blanchet@43314
   413
               entire time available. *)
blanchet@43517
   414
            val actual_slices = get_slices slicing (best_slices ctxt)
blanchet@43314
   415
            val num_actual_slices = length actual_slices
blanchet@43316
   416
            fun monomorphize_facts facts =
blanchet@43316
   417
              let
blanchet@43316
   418
                val repair_context =
blanchet@43316
   419
                  Config.put SMT_Config.verbose debug
blanchet@43316
   420
                  #> Config.put SMT_Config.monomorph_full false
blanchet@43316
   421
                  #> Config.put SMT_Config.monomorph_limit monomorphize_limit
blanchet@43316
   422
                val facts = facts |> map untranslated_fact
blanchet@43316
   423
                (* pseudo-theorem involving the same constants as the subgoal *)
blanchet@43316
   424
                val subgoal_th =
blanchet@43316
   425
                  Logic.list_implies (hyp_ts, concl_t)
blanchet@43316
   426
                  |> Skip_Proof.make_thm thy
blanchet@43316
   427
                val indexed_facts =
blanchet@43316
   428
                  (~1, subgoal_th) :: (0 upto length facts - 1 ~~ map snd facts)
blanchet@43316
   429
              in
blanchet@43316
   430
                SMT_Monomorph.monomorph indexed_facts (repair_context ctxt)
blanchet@43316
   431
                |> fst |> sort (int_ord o pairself fst)
blanchet@43316
   432
                |> filter_out (curry (op =) ~1 o fst)
blanchet@43316
   433
                |> map (Untranslated_Fact o apfst (fst o nth facts))
blanchet@43316
   434
              end
blanchet@43323
   435
            fun run_slice blacklist
blanchet@43321
   436
                          (slice, (time_frac, (complete, default_max_relevant)))
blanchet@43314
   437
                          time_left =
blanchet@38278
   438
              let
blanchet@43314
   439
                val num_facts =
blanchet@43314
   440
                  length facts |> is_none max_relevant
blanchet@43314
   441
                                  ? Integer.min default_max_relevant
blanchet@43509
   442
                val fairly_sound = is_type_sys_fairly_sound type_sys
blanchet@43322
   443
                val facts =
blanchet@43532
   444
                  facts |> not fairly_sound
blanchet@43509
   445
                           ? filter_out (is_dangerous_term o prop_of o snd
blanchet@43509
   446
                                         o untranslated_fact)
blanchet@43509
   447
                        |> take num_facts
blanchet@43323
   448
                        |> not (null blacklist)
blanchet@43323
   449
                           ? filter_out (member (op =) blacklist o fst
blanchet@43323
   450
                                         o untranslated_fact)
blanchet@43460
   451
                        |> polymorphism_of_type_sys type_sys <> Polymorphic
blanchet@43460
   452
                           ? monomorphize_facts
blanchet@43415
   453
                        |> map (atp_translated_fact ctxt)
blanchet@43314
   454
                val real_ms = Real.fromInt o Time.toMilliseconds
blanchet@43314
   455
                val slice_timeout =
blanchet@43314
   456
                  ((real_ms time_left
blanchet@43314
   457
                    |> (if slice < num_actual_slices - 1 then
blanchet@43314
   458
                          curry Real.min (time_frac * real_ms timeout)
blanchet@43314
   459
                        else
blanchet@43314
   460
                          I))
blanchet@43314
   461
                   * 0.001) |> seconds
blanchet@43314
   462
                val _ =
blanchet@43485
   463
                  if debug then
blanchet@43569
   464
                    quote name ^ " slice #" ^ string_of_int (slice + 1) ^
blanchet@43569
   465
                    " with " ^ string_of_int num_facts ^ " fact" ^
blanchet@43569
   466
                    plural_s num_facts ^ " for " ^
blanchet@43569
   467
                    string_from_time slice_timeout ^ "..."
blanchet@43314
   468
                    |> Output.urgent_message
blanchet@43314
   469
                  else
blanchet@43314
   470
                    ()
blanchet@43412
   471
                val (atp_problem, pool, conjecture_offset, facts_offset,
blanchet@43412
   472
                     fact_names) =
blanchet@43580
   473
                  prepare_atp_problem ctxt conj_sym_kind prem_kind type_sys
blanchet@43580
   474
                                      explicit_apply hyp_ts concl_t facts
blanchet@41561
   475
                fun weights () = atp_problem_weights atp_problem
blanchet@41561
   476
                val core =
blanchet@41561
   477
                  File.shell_path command ^ " " ^
blanchet@43517
   478
                  arguments ctxt slice slice_timeout weights ^ " " ^
blanchet@41561
   479
                  File.shell_path prob_file
blanchet@41561
   480
                val command =
blanchet@41561
   481
                  (if measure_run_time then
blanchet@41561
   482
                     "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }"
blanchet@41561
   483
                   else
blanchet@41561
   484
                     "exec " ^ core) ^ " 2>&1"
blanchet@43314
   485
                val _ =
blanchet@43314
   486
                  atp_problem
blanchet@43580
   487
                  |> tptp_strings_for_atp_problem format
blanchet@43314
   488
                  |> cons ("% " ^ command ^ "\n")
blanchet@43314
   489
                  |> File.write_list prob_file
blanchet@43314
   490
                val conjecture_shape =
blanchet@43314
   491
                  conjecture_offset + 1
blanchet@43314
   492
                    upto conjecture_offset + length hyp_ts + 1
blanchet@43314
   493
                  |> map single
blanchet@38278
   494
                val ((output, msecs), res_code) =
blanchet@38278
   495
                  bash_output command
blanchet@38278
   496
                  |>> (if overlord then
blanchet@38278
   497
                         prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
blanchet@38278
   498
                       else
blanchet@38278
   499
                         I)
blanchet@40243
   500
                  |>> (if measure_run_time then split_time else rpair NONE)
blanchet@43320
   501
                val (atp_proof, outcome) =
blanchet@42915
   502
                  extract_tstplike_proof_and_outcome debug verbose complete
blanchet@42915
   503
                      res_code proof_delims known_failures output
blanchet@43320
   504
                  |>> atp_proof_from_tstplike_proof
blanchet@43458
   505
                val (conjecture_shape, facts_offset, fact_names) =
blanchet@43320
   506
                  if is_none outcome then
blanchet@43518
   507
                    repair_conjecture_shape_and_fact_names type_sys output
blanchet@43458
   508
                        conjecture_shape facts_offset fact_names
blanchet@43320
   509
                  else
blanchet@43458
   510
                    (* don't bother repairing *)
blanchet@43458
   511
                    (conjecture_shape, facts_offset, fact_names)
blanchet@43320
   512
                val outcome =
blanchet@43322
   513
                  case outcome of
blanchet@43458
   514
                    NONE =>
blanchet@43518
   515
                    if is_unsound_proof type_sys conjecture_shape facts_offset
blanchet@43518
   516
                                        fact_names atp_proof then
blanchet@43460
   517
                      SOME (UnsoundProof (is_type_sys_virtually_sound type_sys))
blanchet@43458
   518
                    else
blanchet@43458
   519
                      NONE
blanchet@43322
   520
                  | SOME Unprovable =>
blanchet@43323
   521
                    if null blacklist then outcome
blanchet@43322
   522
                    else SOME IncompleteUnprovable
blanchet@43322
   523
                  | _ => outcome
blanchet@43314
   524
              in
blanchet@43412
   525
                ((pool, conjecture_shape, facts_offset, fact_names),
blanchet@43320
   526
                 (output, msecs, atp_proof, outcome))
blanchet@43314
   527
              end
blanchet@38883
   528
            val timer = Timer.startRealTimer ()
blanchet@43323
   529
            fun maybe_run_slice blacklist slice
blanchet@43323
   530
                                (result as (_, (_, msecs0, _, SOME _))) =
blanchet@43323
   531
                let
blanchet@43323
   532
                  val time_left = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@43323
   533
                in
blanchet@43323
   534
                  if Time.<= (time_left, Time.zeroTime) then
blanchet@43323
   535
                    result
blanchet@43323
   536
                  else
blanchet@43323
   537
                    (run_slice blacklist slice time_left
blanchet@43323
   538
                     |> (fn (stuff, (output, msecs, atp_proof, outcome)) =>
blanchet@43323
   539
                            (stuff, (output, int_opt_add msecs0 msecs, atp_proof,
blanchet@43323
   540
                                     outcome))))
blanchet@43323
   541
                end
blanchet@43321
   542
              | maybe_run_slice _ _ result = result
blanchet@43323
   543
            fun maybe_blacklist_facts_and_retry iter blacklist
blanchet@43412
   544
                    (result as ((_, _, facts_offset, fact_names),
blanchet@43458
   545
                                (_, _, atp_proof, SOME (UnsoundProof false)))) =
blanchet@43518
   546
                (case used_facts_in_atp_proof type_sys facts_offset fact_names
blanchet@43412
   547
                                              atp_proof of
blanchet@43323
   548
                   [] => result
blanchet@43323
   549
                 | new_baddies =>
blanchet@43323
   550
                   let val blacklist = new_baddies @ blacklist in
blanchet@43323
   551
                     result
blanchet@43323
   552
                     |> maybe_run_slice blacklist (List.last actual_slices)
blanchet@43323
   553
                     |> iter < atp_blacklist_max_iters
blanchet@43323
   554
                        ? maybe_blacklist_facts_and_retry (iter + 1) blacklist
blanchet@43323
   555
                   end)
blanchet@43323
   556
              | maybe_blacklist_facts_and_retry _ _ result = result
blanchet@43314
   557
          in
blanchet@43412
   558
            ((Symtab.empty, [], 0, Vector.fromList []),
blanchet@43320
   559
             ("", SOME 0, [], SOME InternalError))
blanchet@43322
   560
            |> fold (maybe_run_slice []) actual_slices
blanchet@43322
   561
               (* The ATP found an unsound proof? Automatically try again
blanchet@43322
   562
                  without the offending facts! *)
blanchet@43323
   563
            |> maybe_blacklist_facts_and_retry 0 []
blanchet@43314
   564
          end
blanchet@38278
   565
        else
wenzelm@42815
   566
          error ("Bad executable: " ^ Path.print command ^ ".")
blanchet@38257
   567
blanchet@38257
   568
    (* If the problem file has not been exported, remove it; otherwise, export
blanchet@38257
   569
       the proof file too. *)
blanchet@41561
   570
    fun cleanup prob_file =
blanchet@41561
   571
      if dest_dir = "" then try File.rm prob_file else NONE
blanchet@41561
   572
    fun export prob_file (_, (output, _, _, _)) =
blanchet@40240
   573
      if dest_dir = "" then
blanchet@38257
   574
        ()
blanchet@38257
   575
      else
blanchet@41561
   576
        File.write (Path.explode (Path.implode prob_file ^ "_proof")) output
blanchet@43412
   577
    val ((pool, conjecture_shape, facts_offset, fact_names),
blanchet@43320
   578
         (output, msecs, atp_proof, outcome)) =
blanchet@40242
   579
      with_path cleanup export run_on problem_path_name
blanchet@39732
   580
    val important_message =
blanchet@43480
   581
      if not auto andalso
blanchet@43480
   582
         random_range 0 (atp_important_message_keep_quotient - 1) = 0 then
blanchet@39732
   583
        extract_important_message output
blanchet@39732
   584
      else
blanchet@39732
   585
        ""
blanchet@42613
   586
    fun append_to_message message =
blanchet@42613
   587
      message ^
blanchet@42613
   588
      (if verbose then
blanchet@42613
   589
         "\nATP real CPU time: " ^
blanchet@42613
   590
         string_from_time (Time.fromMilliseconds (the msecs)) ^ "."
blanchet@42613
   591
       else
blanchet@42613
   592
         "") ^
blanchet@42613
   593
      (if important_message <> "" then
blanchet@42613
   594
         "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message
blanchet@42613
   595
       else
blanchet@42613
   596
         "")
blanchet@43518
   597
    val isar_params = (ctxt, debug, isar_shrink_factor, pool, conjecture_shape)
blanchet@42613
   598
    val metis_params =
blanchet@43518
   599
      (proof_banner auto, minimize_command, type_sys, atp_proof, facts_offset,
blanchet@43518
   600
       fact_names, goal, subgoal)
blanchet@42613
   601
    val (outcome, (message, used_facts)) =
blanchet@38257
   602
      case outcome of
blanchet@38257
   603
        NONE =>
blanchet@42613
   604
        (NONE, proof_text isar_proof isar_params metis_params
blanchet@42613
   605
               |>> append_to_message)
blanchet@42613
   606
      | SOME ProofMissing =>
blanchet@42613
   607
        (NONE, metis_proof_text metis_params |>> append_to_message)
blanchet@42615
   608
      | SOME failure => (outcome, (string_for_failure failure, []))
blanchet@38257
   609
  in
blanchet@40445
   610
    {outcome = outcome, message = message, used_facts = used_facts,
blanchet@40242
   611
     run_time_in_msecs = msecs}
blanchet@38257
   612
  end
blanchet@38257
   613
blanchet@40917
   614
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
blanchet@40917
   615
   these are sorted out properly in the SMT module, we have to interpret these
blanchet@40917
   616
   ourselves. *)
blanchet@40932
   617
val remote_smt_failures =
blanchet@40932
   618
  [(22, CantConnect),
blanchet@40932
   619
   (2, NoLibwwwPerl)]
blanchet@41470
   620
val z3_wrapper_failures =
blanchet@41470
   621
  [(10, NoRealZ3),
blanchet@41470
   622
   (11, InternalError),
blanchet@41470
   623
   (12, InternalError),
blanchet@41470
   624
   (13, InternalError)]
blanchet@40932
   625
val z3_failures =
blanchet@41477
   626
  [(101, OutOfResources),
blanchet@41477
   627
   (103, MalformedInput),
blanchet@41470
   628
   (110, MalformedInput)]
blanchet@40932
   629
val unix_failures =
blanchet@40932
   630
  [(139, Crashed)]
blanchet@41470
   631
val smt_failures =
blanchet@42670
   632
  remote_smt_failures @ z3_wrapper_failures @ z3_failures @ unix_failures
blanchet@40797
   633
blanchet@42964
   634
fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
blanchet@42964
   635
    if is_real_cex then Unprovable else IncompleteUnprovable
blanchet@41470
   636
  | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
blanchet@41470
   637
  | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
blanchet@41470
   638
    (case AList.lookup (op =) smt_failures code of
blanchet@40932
   639
       SOME failure => failure
blanchet@41505
   640
     | NONE => UnknownError ("Abnormal termination with exit code " ^
blanchet@41505
   641
                             string_of_int code ^ "."))
blanchet@41470
   642
  | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
blanchet@41470
   643
  | failure_from_smt_failure (SMT_Failure.Other_Failure msg) =
blanchet@42916
   644
    UnknownError msg
wenzelm@28586
   645
blanchet@40946
   646
(* FUDGE *)
blanchet@43517
   647
val smt_max_slices =
blanchet@43517
   648
  Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
blanchet@43517
   649
val smt_slice_fact_frac =
blanchet@43517
   650
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
blanchet@43517
   651
val smt_slice_time_frac =
blanchet@43517
   652
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
blanchet@43517
   653
val smt_slice_min_secs =
blanchet@43517
   654
  Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
blanchet@40636
   655
blanchet@43517
   656
fun smt_filter_loop ctxt name
blanchet@43517
   657
                    ({debug, verbose, overlord, monomorphize_limit, timeout,
blanchet@43517
   658
                      slicing, ...} : params)
blanchet@42612
   659
                    state i smt_filter =
blanchet@40636
   660
  let
blanchet@43517
   661
    val max_slices = if slicing then Config.get ctxt smt_max_slices else 1
blanchet@41483
   662
    val repair_context =
boehmes@41680
   663
      select_smt_solver name
blanchet@41483
   664
      #> Config.put SMT_Config.verbose debug
blanchet@41483
   665
      #> (if overlord then
blanchet@41483
   666
            Config.put SMT_Config.debug_files
blanchet@41483
   667
                       (overlord_file_location_for_prover name
boehmes@41680
   668
                        |> (fn (path, name) => path ^ "/" ^ name))
blanchet@41483
   669
          else
blanchet@41483
   670
            I)
blanchet@43517
   671
      #> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers)
blanchet@43064
   672
      #> Config.put SMT_Config.monomorph_full false
blanchet@43051
   673
      #> Config.put SMT_Config.monomorph_limit monomorphize_limit
blanchet@41483
   674
    val state = state |> Proof.map_context repair_context
blanchet@41483
   675
blanchet@43314
   676
    fun do_slice timeout slice outcome0 time_so_far facts =
blanchet@40795
   677
      let
blanchet@40795
   678
        val timer = Timer.startRealTimer ()
blanchet@40795
   679
        val ms = timeout |> Time.toMilliseconds
blanchet@43314
   680
        val slice_timeout =
blanchet@43314
   681
          if slice < max_slices then
blanchet@41415
   682
            Int.min (ms,
blanchet@43517
   683
                Int.max (1000 * Config.get ctxt smt_slice_min_secs,
blanchet@43517
   684
                    Real.ceil (Config.get ctxt smt_slice_time_frac
blanchet@43517
   685
                               * Real.fromInt ms)))
blanchet@40795
   686
            |> Time.fromMilliseconds
blanchet@40795
   687
          else
blanchet@40795
   688
            timeout
blanchet@40795
   689
        val num_facts = length facts
blanchet@40795
   690
        val _ =
blanchet@43485
   691
          if debug then
blanchet@43485
   692
            quote name ^ " slice " ^ string_of_int slice ^ " with " ^
blanchet@43485
   693
            string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^
blanchet@43485
   694
            string_from_time slice_timeout ^ "..."
blanchet@40795
   695
            |> Output.urgent_message
blanchet@40795
   696
          else
blanchet@40795
   697
            ()
blanchet@41414
   698
        val birth = Timer.checkRealTimer timer
blanchet@41419
   699
        val _ =
blanchet@41459
   700
          if debug then Output.urgent_message "Invoking SMT solver..." else ()
blanchet@41457
   701
        val (outcome, used_facts) =
blanchet@43314
   702
          (case (slice, smt_filter) of
boehmes@41680
   703
             (1, SOME head) => head |> apsnd (apsnd repair_context)
boehmes@41680
   704
           | _ => SMT_Solver.smt_filter_preprocess state facts i)
blanchet@43314
   705
          |> SMT_Solver.smt_filter_apply slice_timeout
blanchet@41480
   706
          |> (fn {outcome, used_facts} => (outcome, used_facts))
blanchet@41457
   707
          handle exn => if Exn.is_interrupt exn then
blanchet@41457
   708
                          reraise exn
blanchet@41457
   709
                        else
blanchet@42916
   710
                          (ML_Compiler.exn_message exn
blanchet@41457
   711
                           |> SMT_Failure.Other_Failure |> SOME, [])
blanchet@41414
   712
        val death = Timer.checkRealTimer timer
blanchet@40795
   713
        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
blanchet@41414
   714
        val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
blanchet@40795
   715
        val too_many_facts_perhaps =
blanchet@40795
   716
          case outcome of
blanchet@40795
   717
            NONE => false
blanchet@40795
   718
          | SOME (SMT_Failure.Counterexample _) => false
blanchet@43314
   719
          | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
blanchet@43485
   720
          | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
blanchet@40795
   721
          | SOME SMT_Failure.Out_Of_Memory => true
blanchet@41459
   722
          | SOME (SMT_Failure.Other_Failure _) => true
blanchet@40795
   723
        val timeout = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@40795
   724
      in
blanchet@43314
   725
        if too_many_facts_perhaps andalso slice < max_slices andalso
blanchet@40795
   726
           num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
blanchet@41415
   727
          let
blanchet@43485
   728
            val new_num_facts =
blanchet@43517
   729
              Real.ceil (Config.get ctxt smt_slice_fact_frac
blanchet@43517
   730
                         * Real.fromInt num_facts)
blanchet@43485
   731
            val _ =
blanchet@43485
   732
              if verbose andalso is_some outcome then
blanchet@43485
   733
                quote name ^ " invoked with " ^ string_of_int num_facts ^
blanchet@43485
   734
                " fact" ^ plural_s num_facts ^ ": " ^
blanchet@43485
   735
                string_for_failure (failure_from_smt_failure (the outcome)) ^
blanchet@43485
   736
                " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^
blanchet@43509
   737
                plural_s new_num_facts ^ "..."
blanchet@43485
   738
                |> Output.urgent_message
blanchet@43485
   739
              else
blanchet@43485
   740
                ()
blanchet@43314
   741
          in
blanchet@43485
   742
            facts |> take new_num_facts
blanchet@43485
   743
                  |> do_slice timeout (slice + 1) outcome0 time_so_far
blanchet@43314
   744
          end
blanchet@40795
   745
        else
blanchet@40795
   746
          {outcome = if is_none outcome then NONE else the outcome0,
blanchet@41414
   747
           used_facts = used_facts,
blanchet@41414
   748
           run_time_in_msecs = SOME (Time.toMilliseconds time_so_far)}
blanchet@40636
   749
      end
blanchet@43314
   750
  in do_slice timeout 1 NONE Time.zeroTime end
blanchet@40636
   751
blanchet@40914
   752
(* taken from "Mirabelle" and generalized *)
blanchet@40914
   753
fun can_apply timeout tac state i =
blanchet@40914
   754
  let
blanchet@40914
   755
    val {context = ctxt, facts, goal} = Proof.goal state
blanchet@40914
   756
    val full_tac = Method.insert_tac facts i THEN tac ctxt i
blanchet@40914
   757
  in
blanchet@40914
   758
    case try (TimeLimit.timeLimit timeout (Seq.pull o full_tac)) goal of
blanchet@40914
   759
      SOME (SOME _) => true
blanchet@40914
   760
    | _ => false
blanchet@40914
   761
  end
blanchet@40914
   762
blanchet@41400
   763
val smt_metis_timeout = seconds 1.0
blanchet@40914
   764
blanchet@40941
   765
fun can_apply_metis debug state i ths =
blanchet@40941
   766
  can_apply smt_metis_timeout
blanchet@40941
   767
            (Config.put Metis_Tactics.verbose debug
blanchet@40941
   768
             #> (fn ctxt => Metis_Tactics.metis_tac ctxt ths)) state i
blanchet@40914
   769
blanchet@43315
   770
fun run_smt_solver auto name (params as {debug, verbose, ...}) minimize_command
blanchet@43314
   771
                   ({state, subgoal, subgoal_count, facts, smt_filter, ...}
blanchet@43314
   772
                    : prover_problem) =
blanchet@36379
   773
  let
blanchet@41483
   774
    val ctxt = Proof.context_of state
blanchet@41502
   775
    val num_facts = length facts
blanchet@41502
   776
    val facts = facts ~~ (0 upto num_facts - 1)
blanchet@43517
   777
                |> map (smt_weighted_fact ctxt num_facts)
blanchet@40422
   778
    val {outcome, used_facts, run_time_in_msecs} =
blanchet@43517
   779
      smt_filter_loop ctxt name params state subgoal smt_filter facts
blanchet@40966
   780
    val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
blanchet@41470
   781
    val outcome = outcome |> Option.map failure_from_smt_failure
blanchet@40244
   782
    val message =
blanchet@40425
   783
      case outcome of
blanchet@40425
   784
        NONE =>
blanchet@40914
   785
        let
blanchet@41399
   786
          val (method, settings) =
blanchet@40941
   787
            if can_apply_metis debug state subgoal (map snd used_facts) then
blanchet@41399
   788
              ("metis", "")
blanchet@40941
   789
            else
boehmes@41680
   790
              ("smt", if name = SMT_Solver.solver_name_of ctxt then ""
boehmes@41680
   791
                      else "smt_solver = " ^ maybe_quote name)
blanchet@40914
   792
        in
blanchet@40914
   793
          try_command_line (proof_banner auto)
blanchet@41399
   794
              (apply_on_subgoal settings subgoal subgoal_count ^
blanchet@41399
   795
               command_call method (map fst other_lemmas)) ^
blanchet@40966
   796
          minimize_line minimize_command
blanchet@41414
   797
                        (map fst (other_lemmas @ chained_lemmas)) ^
blanchet@41414
   798
          (if verbose then
blanchet@41414
   799
             "\nSMT solver real CPU time: " ^
blanchet@41414
   800
             string_from_time (Time.fromMilliseconds (the run_time_in_msecs)) ^
blanchet@41414
   801
             "."
blanchet@41414
   802
           else
blanchet@41414
   803
             "")
blanchet@40914
   804
        end
blanchet@42615
   805
      | SOME failure => string_for_failure failure
blanchet@40244
   806
  in
blanchet@40914
   807
    {outcome = outcome, used_facts = map fst used_facts,
blanchet@40445
   808
     run_time_in_msecs = run_time_in_msecs, message = message}
blanchet@40244
   809
  end
blanchet@40244
   810
blanchet@43315
   811
fun get_prover ctxt auto name =
wenzelm@43232
   812
  let val thy = Proof_Context.theory_of ctxt in
blanchet@41189
   813
    if is_smt_prover ctxt name then
blanchet@43315
   814
      run_smt_solver auto name
blanchet@42591
   815
    else if member (op =) (supported_atps thy) name then
blanchet@43315
   816
      run_atp auto name (get_atp thy name)
blanchet@41189
   817
    else
blanchet@41189
   818
      error ("No such prover: " ^ name ^ ".")
blanchet@41189
   819
  end
blanchet@40244
   820
wenzelm@28582
   821
end;