src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49334 340187063d84
parent 49329 ee33ba3c0e05
child 49336 c552d7f1720b
permissions -rw-r--r--
use async manager to manage MaSh learners to make sure they get killed cleanly
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@47168
    12
  type stature = ATP_Problem_Generate.stature
blanchet@47148
    13
  type type_enc = ATP_Problem_Generate.type_enc
blanchet@47148
    14
  type reconstructor = ATP_Proof_Reconstruct.reconstructor
blanchet@47148
    15
  type play = ATP_Proof_Reconstruct.play
blanchet@47148
    16
  type minimize_command = ATP_Proof_Reconstruct.minimize_command
blanchet@39733
    17
blanchet@46445
    18
  datatype mode = Auto_Try | Try | Normal | Auto_Minimize | Minimize
blanchet@43862
    19
blanchet@35967
    20
  type params =
blanchet@41456
    21
    {debug: bool,
blanchet@35967
    22
     verbose: bool,
blanchet@36141
    23
     overlord: bool,
blanchet@41456
    24
     blocking: bool,
blanchet@40240
    25
     provers: string list,
blanchet@45256
    26
     type_enc: string option,
blanchet@47129
    27
     strict: bool,
blanchet@46391
    28
     lam_trans: string option,
blanchet@47237
    29
     uncurried_aliases: bool option,
blanchet@49329
    30
     fact_filter: string option,
blanchet@49329
    31
     max_facts: int option,
blanchet@49308
    32
     fact_thresholds: real * real,
blanchet@48977
    33
     max_mono_iters: int option,
blanchet@48977
    34
     max_new_mono_instances: int option,
blanchet@35967
    35
     isar_proof: bool,
blanchet@36916
    36
     isar_shrink_factor: int,
blanchet@46577
    37
     slice: bool,
blanchet@46578
    38
     minimize: bool option,
blanchet@39229
    39
     timeout: Time.time,
blanchet@43856
    40
     preplay_timeout: Time.time,
blanchet@39229
    41
     expect: string}
blanchet@39733
    42
blanchet@49303
    43
  type relevance_fudge =
blanchet@49303
    44
    {local_const_multiplier : real,
blanchet@49303
    45
     worse_irrel_freq : real,
blanchet@49303
    46
     higher_order_irrel_weight : real,
blanchet@49303
    47
     abs_rel_weight : real,
blanchet@49303
    48
     abs_irrel_weight : real,
blanchet@49303
    49
     skolem_irrel_weight : real,
blanchet@49303
    50
     theory_const_rel_weight : real,
blanchet@49303
    51
     theory_const_irrel_weight : real,
blanchet@49303
    52
     chained_const_irrel_weight : real,
blanchet@49303
    53
     intro_bonus : real,
blanchet@49303
    54
     elim_bonus : real,
blanchet@49303
    55
     simp_bonus : real,
blanchet@49303
    56
     local_bonus : real,
blanchet@49303
    57
     assum_bonus : real,
blanchet@49303
    58
     chained_bonus : real,
blanchet@49303
    59
     max_imperfect : real,
blanchet@49303
    60
     max_imperfect_exp : real,
blanchet@49303
    61
     threshold_divisor : real,
blanchet@49303
    62
     ridiculous_threshold : real}
blanchet@49303
    63
blanchet@41338
    64
  datatype prover_fact =
blanchet@47168
    65
    Untranslated_Fact of (string * stature) * thm |
blanchet@47168
    66
    SMT_Weighted_Fact of (string * stature) * (int option * thm)
blanchet@40242
    67
blanchet@40242
    68
  type prover_problem =
blanchet@39564
    69
    {state: Proof.state,
blanchet@39242
    70
     goal: thm,
blanchet@39242
    71
     subgoal: int,
blanchet@40246
    72
     subgoal_count: int,
blanchet@48402
    73
     facts: prover_fact list}
blanchet@39733
    74
blanchet@40242
    75
  type prover_result =
blanchet@36370
    76
    {outcome: failure option,
blanchet@47168
    77
     used_facts: (string * stature) list,
blanchet@46231
    78
     run_time: Time.time,
blanchet@43893
    79
     preplay: unit -> play,
blanchet@44102
    80
     message: play -> string,
blanchet@44102
    81
     message_tail: string}
blanchet@39733
    82
blanchet@43892
    83
  type prover =
blanchet@46391
    84
    params -> ((string * string list) list -> string -> minimize_command)
blanchet@46391
    85
    -> prover_problem -> prover_result
blanchet@35867
    86
blanchet@43933
    87
  val dest_dir : string Config.T
blanchet@43933
    88
  val problem_prefix : string Config.T
blanchet@49158
    89
  val completish : bool Config.T
blanchet@45456
    90
  val atp_full_names : bool Config.T
blanchet@43517
    91
  val smt_triggers : bool Config.T
blanchet@43517
    92
  val smt_weights : bool Config.T
blanchet@43517
    93
  val smt_weight_min_facts : int Config.T
blanchet@43517
    94
  val smt_min_weight : int Config.T
blanchet@43517
    95
  val smt_max_weight : int Config.T
blanchet@43517
    96
  val smt_max_weight_index : int Config.T
blanchet@41502
    97
  val smt_weight_curve : (int -> int) Unsynchronized.ref
blanchet@43517
    98
  val smt_max_slices : int Config.T
blanchet@43517
    99
  val smt_slice_fact_frac : real Config.T
blanchet@43517
   100
  val smt_slice_time_frac : real Config.T
blanchet@43517
   101
  val smt_slice_min_secs : int Config.T
blanchet@49334
   102
  val SledgehammerN : string
blanchet@46390
   103
  val plain_metis : reconstructor
blanchet@41483
   104
  val select_smt_solver : string -> Proof.context -> Proof.context
blanchet@46391
   105
  val extract_reconstructor :
blanchet@46432
   106
    params -> reconstructor -> string * (string * string list) list
blanchet@46250
   107
  val is_reconstructor : string -> bool
blanchet@43891
   108
  val is_atp : theory -> string -> bool
blanchet@41335
   109
  val is_smt_prover : Proof.context -> string -> bool
blanchet@48977
   110
  val is_ho_atp: Proof.context -> string -> bool
blanchet@43785
   111
  val is_unit_equational_atp : Proof.context -> string -> bool
blanchet@42591
   112
  val is_prover_supported : Proof.context -> string -> bool
blanchet@40253
   113
  val is_prover_installed : Proof.context -> string -> bool
blanchet@49308
   114
  val default_max_facts_for_prover : Proof.context -> bool -> string -> int
blanchet@43785
   115
  val is_unit_equality : term -> bool
blanchet@43793
   116
  val is_appropriate_prop_for_prover : Proof.context -> string -> term -> bool
blanchet@40615
   117
  val is_built_in_const_for_prover :
blanchet@41584
   118
    Proof.context -> string -> string * typ -> term list -> bool * term list
blanchet@41335
   119
  val atp_relevance_fudge : relevance_fudge
blanchet@41335
   120
  val smt_relevance_fudge : relevance_fudge
blanchet@41189
   121
  val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
blanchet@41502
   122
  val weight_smt_fact :
blanchet@47168
   123
    Proof.context -> int -> ((string * stature) * thm) * int
blanchet@47168
   124
    -> (string * stature) * (int option * thm)
blanchet@47168
   125
  val untranslated_fact : prover_fact -> (string * stature) * thm
blanchet@41483
   126
  val smt_weighted_fact :
blanchet@43517
   127
    Proof.context -> int -> prover_fact * int
blanchet@47168
   128
    -> (string * stature) * (int option * thm)
blanchet@42591
   129
  val supported_provers : Proof.context -> unit
blanchet@40240
   130
  val kill_provers : unit -> unit
blanchet@40240
   131
  val running_provers : unit -> unit
blanchet@40240
   132
  val messages : int option -> unit
blanchet@43874
   133
  val filter_used_facts : ''a list -> (''a * 'b) list -> (''a * 'b) list
blanchet@43862
   134
  val get_prover : Proof.context -> mode -> string -> prover
wenzelm@28477
   135
end;
wenzelm@28477
   136
blanchet@41335
   137
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
wenzelm@28477
   138
struct
wenzelm@28477
   139
blanchet@43926
   140
open ATP_Util
blanchet@38262
   141
open ATP_Problem
blanchet@39731
   142
open ATP_Proof
blanchet@38262
   143
open ATP_Systems
blanchet@47148
   144
open ATP_Problem_Generate
blanchet@47148
   145
open ATP_Proof_Reconstruct
blanchet@46392
   146
open Metis_Tactic
blanchet@38257
   147
open Sledgehammer_Util
blanchet@49303
   148
blanchet@37583
   149
blanchet@37583
   150
(** The Sledgehammer **)
blanchet@37583
   151
blanchet@46445
   152
datatype mode = Auto_Try | Try | Normal | Auto_Minimize | Minimize
blanchet@43862
   153
blanchet@46247
   154
(* Identifier that distinguishes Sledgehammer from other tools that could use
blanchet@38348
   155
   "Async_Manager". *)
blanchet@49334
   156
val SledgehammerN = "Sledgehammer"
blanchet@37585
   157
blanchet@46391
   158
val reconstructor_names = [metisN, smtN]
blanchet@47193
   159
val plain_metis = Metis (hd partial_type_encs, combsN)
blanchet@46432
   160
val is_reconstructor = member (op =) reconstructor_names
blanchet@44069
   161
blanchet@43891
   162
val is_atp = member (op =) o supported_atps
blanchet@43891
   163
blanchet@44074
   164
val select_smt_solver = Context.proof_map o SMT_Config.select_solver
blanchet@41483
   165
blanchet@46247
   166
fun is_smt_prover ctxt = member (op =) (SMT_Solver.available_solvers_of ctxt)
blanchet@40243
   167
blanchet@45461
   168
fun is_atp_for_format is_format ctxt name =
blanchet@43785
   169
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43785
   170
    case try (get_atp thy) name of
blanchet@48469
   171
      SOME config =>
blanchet@47237
   172
      exists (fn (_, (_, ((_, format, _, _, _), _))) => is_format format)
blanchet@48469
   173
             (#best_slices (config ()) ctxt)
blanchet@43785
   174
    | NONE => false
blanchet@43785
   175
  end
blanchet@43785
   176
blanchet@45461
   177
val is_unit_equational_atp = is_atp_for_format (curry (op =) CNF_UEQ)
blanchet@46174
   178
val is_ho_atp = is_atp_for_format is_format_higher_order
blanchet@45461
   179
blanchet@46247
   180
fun is_prover_supported ctxt =
wenzelm@43232
   181
  let val thy = Proof_Context.theory_of ctxt in
blanchet@46250
   182
    is_reconstructor orf is_atp thy orf is_smt_prover ctxt
blanchet@41189
   183
  end
blanchet@40253
   184
boehmes@41680
   185
fun is_prover_installed ctxt =
blanchet@46250
   186
  is_reconstructor orf is_smt_prover ctxt orf
blanchet@43891
   187
  is_atp_installed (Proof_Context.theory_of ctxt)
blanchet@41189
   188
blanchet@46577
   189
fun get_slices slice slices =
blanchet@46577
   190
  (0 upto length slices - 1) ~~ slices |> not slice ? (List.last #> single)
blanchet@43314
   191
blanchet@49308
   192
val reconstructor_default_max_facts = 20
blanchet@43891
   193
blanchet@49308
   194
fun default_max_facts_for_prover ctxt slice name =
wenzelm@43232
   195
  let val thy = Proof_Context.theory_of ctxt in
blanchet@46250
   196
    if is_reconstructor name then
blanchet@49308
   197
      reconstructor_default_max_facts
blanchet@43891
   198
    else if is_atp thy name then
blanchet@47235
   199
      fold (Integer.max o #1 o fst o snd o snd o snd)
blanchet@48469
   200
           (get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0
blanchet@43891
   201
    else (* is_smt_prover ctxt name *)
blanchet@43891
   202
      SMT_Solver.default_max_relevant ctxt name
blanchet@41189
   203
  end
blanchet@40244
   204
blanchet@43797
   205
fun is_if (@{const_name If}, _) = true
blanchet@43797
   206
  | is_if _ = false
blanchet@43797
   207
blanchet@43797
   208
(* Beware of "if and only if" (which is translated as such) and "If" (which is
blanchet@43797
   209
   translated to conditional equations). *)
blanchet@43797
   210
fun is_good_unit_equality T t u =
blanchet@43797
   211
  T <> @{typ bool} andalso not (exists (exists_Const is_if) [t, u])
blanchet@43797
   212
blanchet@43785
   213
fun is_unit_equality (@{const Trueprop} $ t) = is_unit_equality t
blanchet@43785
   214
  | is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) =
blanchet@43785
   215
    is_unit_equality t
blanchet@43785
   216
  | is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) =
blanchet@43785
   217
    is_unit_equality t
blanchet@43797
   218
  | is_unit_equality (Const (@{const_name "=="}, Type (_, [T, _])) $ t $ u) =
blanchet@43797
   219
    is_good_unit_equality T t u
blanchet@43797
   220
  | is_unit_equality (Const (@{const_name HOL.eq}, Type (_ , [T, _])) $ t $ u) =
blanchet@43797
   221
    is_good_unit_equality T t u
blanchet@43785
   222
  | is_unit_equality _ = false
blanchet@43785
   223
blanchet@43793
   224
fun is_appropriate_prop_for_prover ctxt name =
blanchet@43785
   225
  if is_unit_equational_atp ctxt name then is_unit_equality else K true
blanchet@43785
   226
blanchet@41483
   227
fun is_built_in_const_for_prover ctxt name =
blanchet@41483
   228
  if is_smt_prover ctxt name then
blanchet@41584
   229
    let val ctxt = ctxt |> select_smt_solver name in
blanchet@41584
   230
      fn x => fn ts =>
blanchet@41584
   231
         if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
blanchet@41584
   232
           (true, [])
blanchet@41584
   233
         else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
blanchet@41584
   234
           (true, ts)
blanchet@41584
   235
         else
blanchet@41584
   236
           (false, ts)
blanchet@41584
   237
    end
blanchet@41483
   238
  else
blanchet@41584
   239
    fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts)
blanchet@40252
   240
blanchet@40251
   241
(* FUDGE *)
blanchet@40251
   242
val atp_relevance_fudge =
blanchet@43603
   243
  {local_const_multiplier = 1.5,
blanchet@41407
   244
   worse_irrel_freq = 100.0,
blanchet@40251
   245
   higher_order_irrel_weight = 1.05,
blanchet@40251
   246
   abs_rel_weight = 0.5,
blanchet@40251
   247
   abs_irrel_weight = 2.0,
blanchet@48949
   248
   skolem_irrel_weight = 0.05,
blanchet@40251
   249
   theory_const_rel_weight = 0.5,
blanchet@40251
   250
   theory_const_irrel_weight = 0.25,
blanchet@43600
   251
   chained_const_irrel_weight = 0.25,
blanchet@40251
   252
   intro_bonus = 0.15,
blanchet@40251
   253
   elim_bonus = 0.15,
blanchet@40251
   254
   simp_bonus = 0.15,
blanchet@40251
   255
   local_bonus = 0.55,
blanchet@40251
   256
   assum_bonus = 1.05,
blanchet@40251
   257
   chained_bonus = 1.5,
blanchet@40251
   258
   max_imperfect = 11.5,
blanchet@40251
   259
   max_imperfect_exp = 1.0,
blanchet@40251
   260
   threshold_divisor = 2.0,
blanchet@41341
   261
   ridiculous_threshold = 0.01}
blanchet@40251
   262
blanchet@40252
   263
(* FUDGE (FIXME) *)
blanchet@40251
   264
val smt_relevance_fudge =
blanchet@43603
   265
  {local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
blanchet@41407
   266
   worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
blanchet@40252
   267
   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
blanchet@40252
   268
   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
blanchet@40252
   269
   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
blanchet@40252
   270
   skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
blanchet@40252
   271
   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
blanchet@40252
   272
   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
blanchet@43600
   273
   chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge,
blanchet@40252
   274
   intro_bonus = #intro_bonus atp_relevance_fudge,
blanchet@40252
   275
   elim_bonus = #elim_bonus atp_relevance_fudge,
blanchet@40252
   276
   simp_bonus = #simp_bonus atp_relevance_fudge,
blanchet@40252
   277
   local_bonus = #local_bonus atp_relevance_fudge,
blanchet@40252
   278
   assum_bonus = #assum_bonus atp_relevance_fudge,
blanchet@40252
   279
   chained_bonus = #chained_bonus atp_relevance_fudge,
blanchet@40252
   280
   max_imperfect = #max_imperfect atp_relevance_fudge,
blanchet@40252
   281
   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
blanchet@40252
   282
   threshold_divisor = #threshold_divisor atp_relevance_fudge,
blanchet@40252
   283
   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
blanchet@40251
   284
blanchet@41189
   285
fun relevance_fudge_for_prover ctxt name =
blanchet@41189
   286
  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
blanchet@40251
   287
blanchet@42591
   288
fun supported_provers ctxt =
blanchet@40241
   289
  let
wenzelm@43232
   290
    val thy = Proof_Context.theory_of ctxt
blanchet@40241
   291
    val (remote_provers, local_provers) =
blanchet@46391
   292
      reconstructor_names @
blanchet@42591
   293
      sort_strings (supported_atps thy) @
blanchet@42591
   294
      sort_strings (SMT_Solver.available_solvers_of ctxt)
blanchet@40241
   295
      |> List.partition (String.isPrefix remote_prefix)
blanchet@40241
   296
  in
blanchet@42591
   297
    Output.urgent_message ("Supported provers: " ^
blanchet@40446
   298
                           commas (local_provers @ remote_provers) ^ ".")
blanchet@40241
   299
  end
blanchet@35967
   300
blanchet@49334
   301
fun kill_provers () = Async_Manager.kill_threads SledgehammerN "prover"
blanchet@49334
   302
fun running_provers () = Async_Manager.running_threads SledgehammerN "prover"
blanchet@49334
   303
val messages = Async_Manager.thread_messages SledgehammerN "prover"
blanchet@40240
   304
blanchet@49303
   305
blanchet@40240
   306
(** problems, results, ATPs, etc. **)
blanchet@35967
   307
blanchet@35967
   308
type params =
blanchet@41456
   309
  {debug: bool,
blanchet@35967
   310
   verbose: bool,
blanchet@36141
   311
   overlord: bool,
blanchet@41456
   312
   blocking: bool,
blanchet@40240
   313
   provers: string list,
blanchet@45256
   314
   type_enc: string option,
blanchet@47129
   315
   strict: bool,
blanchet@46391
   316
   lam_trans: string option,
blanchet@47237
   317
   uncurried_aliases: bool option,
blanchet@49329
   318
   fact_filter: string option,
blanchet@49329
   319
   max_facts: int option,
blanchet@49308
   320
   fact_thresholds: real * real,
blanchet@48977
   321
   max_mono_iters: int option,
blanchet@48977
   322
   max_new_mono_instances: int option,
blanchet@35967
   323
   isar_proof: bool,
blanchet@36916
   324
   isar_shrink_factor: int,
blanchet@46577
   325
   slice: bool,
blanchet@46578
   326
   minimize: bool option,
blanchet@39229
   327
   timeout: Time.time,
blanchet@43856
   328
   preplay_timeout: Time.time,
blanchet@39229
   329
   expect: string}
blanchet@35867
   330
blanchet@49303
   331
type relevance_fudge =
blanchet@49303
   332
  {local_const_multiplier : real,
blanchet@49303
   333
   worse_irrel_freq : real,
blanchet@49303
   334
   higher_order_irrel_weight : real,
blanchet@49303
   335
   abs_rel_weight : real,
blanchet@49303
   336
   abs_irrel_weight : real,
blanchet@49303
   337
   skolem_irrel_weight : real,
blanchet@49303
   338
   theory_const_rel_weight : real,
blanchet@49303
   339
   theory_const_irrel_weight : real,
blanchet@49303
   340
   chained_const_irrel_weight : real,
blanchet@49303
   341
   intro_bonus : real,
blanchet@49303
   342
   elim_bonus : real,
blanchet@49303
   343
   simp_bonus : real,
blanchet@49303
   344
   local_bonus : real,
blanchet@49303
   345
   assum_bonus : real,
blanchet@49303
   346
   chained_bonus : real,
blanchet@49303
   347
   max_imperfect : real,
blanchet@49303
   348
   max_imperfect_exp : real,
blanchet@49303
   349
   threshold_divisor : real,
blanchet@49303
   350
   ridiculous_threshold : real}
blanchet@49303
   351
blanchet@41338
   352
datatype prover_fact =
blanchet@47168
   353
  Untranslated_Fact of (string * stature) * thm |
blanchet@47168
   354
  SMT_Weighted_Fact of (string * stature) * (int option * thm)
blanchet@40242
   355
blanchet@40242
   356
type prover_problem =
blanchet@39564
   357
  {state: Proof.state,
blanchet@39242
   358
   goal: thm,
blanchet@39242
   359
   subgoal: int,
blanchet@40246
   360
   subgoal_count: int,
blanchet@48402
   361
   facts: prover_fact list}
blanchet@35867
   362
blanchet@40242
   363
type prover_result =
blanchet@36370
   364
  {outcome: failure option,
blanchet@47168
   365
   used_facts: (string * stature) list,
blanchet@46231
   366
   run_time: Time.time,
blanchet@43893
   367
   preplay: unit -> play,
blanchet@44102
   368
   message: play -> string,
blanchet@44102
   369
   message_tail: string}
blanchet@35867
   370
blanchet@43892
   371
type prover =
blanchet@46391
   372
  params -> ((string * string list) list -> string -> minimize_command)
blanchet@46391
   373
  -> prover_problem -> prover_result
blanchet@35867
   374
blanchet@38257
   375
(* configuration attributes *)
wenzelm@28484
   376
blanchet@43933
   377
(* Empty string means create files in Isabelle's temporary files directory. *)
wenzelm@43487
   378
val dest_dir =
wenzelm@43487
   379
  Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
wenzelm@43487
   380
val problem_prefix =
wenzelm@43487
   381
  Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
blanchet@49158
   382
val completish =
blanchet@49158
   383
  Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
wenzelm@32995
   384
blanchet@43933
   385
(* In addition to being easier to read, readable names are often much shorter,
blanchet@45253
   386
   especially if types are mangled in names. This makes a difference for some
blanchet@45253
   387
   provers (e.g., E). For these reason, short names are enabled by default. *)
blanchet@45456
   388
val atp_full_names =
blanchet@45456
   389
  Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
blanchet@43933
   390
blanchet@43517
   391
val smt_triggers =
blanchet@43517
   392
  Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
blanchet@43517
   393
val smt_weights =
blanchet@43517
   394
  Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
blanchet@43517
   395
val smt_weight_min_facts =
blanchet@43517
   396
  Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
blanchet@41502
   397
blanchet@41502
   398
(* FUDGE *)
blanchet@43517
   399
val smt_min_weight =
blanchet@43517
   400
  Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
blanchet@43517
   401
val smt_max_weight =
blanchet@43517
   402
  Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
blanchet@43517
   403
val smt_max_weight_index =
blanchet@43517
   404
  Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
blanchet@41502
   405
val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
blanchet@41502
   406
blanchet@43517
   407
fun smt_fact_weight ctxt j num_facts =
blanchet@43517
   408
  if Config.get ctxt smt_weights andalso
blanchet@43517
   409
     num_facts >= Config.get ctxt smt_weight_min_facts then
blanchet@43517
   410
    let
blanchet@43517
   411
      val min = Config.get ctxt smt_min_weight
blanchet@43517
   412
      val max = Config.get ctxt smt_max_weight
blanchet@43517
   413
      val max_index = Config.get ctxt smt_max_weight_index
blanchet@43517
   414
      val curve = !smt_weight_curve
blanchet@43517
   415
    in
blanchet@43517
   416
      SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
blanchet@43517
   417
            div curve max_index)
blanchet@43517
   418
    end
blanchet@41502
   419
  else
blanchet@41502
   420
    NONE
blanchet@41502
   421
blanchet@43517
   422
fun weight_smt_fact ctxt num_facts ((info, th), j) =
blanchet@43517
   423
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43517
   424
    (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
blanchet@43517
   425
  end
blanchet@38257
   426
blanchet@41339
   427
fun untranslated_fact (Untranslated_Fact p) = p
blanchet@41483
   428
  | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
blanchet@41502
   429
fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
blanchet@43517
   430
  | smt_weighted_fact ctxt num_facts (fact, j) =
blanchet@43517
   431
    (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts
blanchet@41502
   432
blanchet@41561
   433
fun overlord_file_location_for_prover prover =
blanchet@41561
   434
  (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
blanchet@41561
   435
blanchet@43874
   436
fun with_path cleanup after f path =
blanchet@43874
   437
  Exn.capture f path
blanchet@43874
   438
  |> tap (fn _ => cleanup path)
blanchet@43874
   439
  |> Exn.release
blanchet@43874
   440
  |> tap (after path)
blanchet@43874
   441
blanchet@43893
   442
fun proof_banner mode name =
blanchet@43874
   443
  case mode of
blanchet@43874
   444
    Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
blanchet@43874
   445
  | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
blanchet@43893
   446
  | _ => "Try this"
blanchet@43874
   447
blanchet@46432
   448
fun bunch_of_reconstructors needs_full_types lam_trans =
blanchet@47233
   449
  [(false, Metis (partial_type_enc, lam_trans false)),
blanchet@47124
   450
   (true, Metis (full_type_enc, lam_trans false)),
blanchet@47233
   451
   (false, Metis (no_typesN, lam_trans true)),
blanchet@47233
   452
   (true, Metis (really_full_type_enc, lam_trans true)),
blanchet@46432
   453
   (true, SMT)]
blanchet@46432
   454
  |> map_filter (fn (full_types, reconstr) =>
blanchet@46432
   455
                    if needs_full_types andalso not full_types then NONE
blanchet@46432
   456
                    else SOME reconstr)
blanchet@46432
   457
blanchet@46432
   458
fun extract_reconstructor ({type_enc, lam_trans, ...} : params)
blanchet@46432
   459
                          (Metis (type_enc', lam_trans')) =
blanchet@46432
   460
    let
blanchet@46432
   461
      val override_params =
blanchet@46432
   462
        (if is_none type_enc andalso type_enc' = hd partial_type_encs then
blanchet@46432
   463
           []
blanchet@46432
   464
         else
blanchet@46437
   465
           [("type_enc", [hd (unalias_type_enc type_enc')])]) @
blanchet@46432
   466
        (if is_none lam_trans andalso lam_trans' = metis_default_lam_trans then
blanchet@46432
   467
           []
blanchet@46432
   468
         else
blanchet@46432
   469
           [("lam_trans", [lam_trans'])])
blanchet@46432
   470
    in (metisN, override_params) end
blanchet@46432
   471
  | extract_reconstructor _ SMT = (smtN, [])
blanchet@46432
   472
blanchet@43874
   473
(* based on "Mirabelle.can_apply" and generalized *)
blanchet@43875
   474
fun timed_apply timeout tac state i =
blanchet@43874
   475
  let
blanchet@43874
   476
    val {context = ctxt, facts, goal} = Proof.goal state
blanchet@43874
   477
    val full_tac = Method.insert_tac facts i THEN tac ctxt i
blanchet@43875
   478
  in TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal end
blanchet@43874
   479
blanchet@46390
   480
fun tac_for_reconstructor (Metis (type_enc, lam_trans)) =
blanchet@46392
   481
    metis_tac [type_enc] lam_trans
blanchet@46250
   482
  | tac_for_reconstructor SMT = SMT_Solver.smt_tac
blanchet@43875
   483
blanchet@46391
   484
fun timed_reconstructor reconstr debug timeout ths =
blanchet@45515
   485
  (Config.put Metis_Tactic.verbose debug
blanchet@46428
   486
   #> Config.put SMT_Config.verbose debug
blanchet@46391
   487
   #> (fn ctxt => tac_for_reconstructor reconstr ctxt ths))
blanchet@43875
   488
  |> timed_apply timeout
blanchet@43874
   489
blanchet@43874
   490
fun filter_used_facts used = filter (member (op =) used o fst)
blanchet@43874
   491
blanchet@46652
   492
fun play_one_line_proof mode debug verbose timeout pairs state i preferred
blanchet@46391
   493
                        reconstrs =
blanchet@43875
   494
  let
blanchet@46250
   495
    val _ =
blanchet@46250
   496
      if mode = Minimize andalso Time.> (timeout, Time.zeroTime) then
blanchet@46250
   497
        Output.urgent_message "Preplaying proof..."
blanchet@46250
   498
      else
blanchet@46250
   499
        ()
blanchet@46652
   500
    val ths = pairs |> sort_wrt (fst o fst) |> map snd
blanchet@46391
   501
    fun get_preferred reconstrs =
blanchet@46391
   502
      if member (op =) reconstrs preferred then preferred
blanchet@46391
   503
      else List.last reconstrs
blanchet@46391
   504
    fun play [] [] = Failed_to_Play (get_preferred reconstrs)
blanchet@46250
   505
      | play timed_outs [] =
blanchet@46250
   506
        Trust_Playable (get_preferred timed_outs, SOME timeout)
blanchet@46391
   507
      | play timed_out (reconstr :: reconstrs) =
blanchet@46249
   508
        let
blanchet@46249
   509
          val _ =
blanchet@46249
   510
            if verbose then
blanchet@46425
   511
              "Trying \"" ^ string_for_reconstructor reconstr ^ "\" for " ^
blanchet@46249
   512
              string_from_time timeout ^ "..."
blanchet@46249
   513
              |> Output.urgent_message
blanchet@46249
   514
            else
blanchet@46249
   515
              ()
blanchet@46249
   516
          val timer = Timer.startRealTimer ()
blanchet@46249
   517
        in
blanchet@46391
   518
          case timed_reconstructor reconstr debug timeout ths state i of
blanchet@46391
   519
            SOME (SOME _) => Played (reconstr, Timer.checkRealTimer timer)
blanchet@46391
   520
          | _ => play timed_out reconstrs
blanchet@43885
   521
        end
blanchet@46391
   522
        handle TimeLimit.TimeOut => play (reconstr :: timed_out) reconstrs
blanchet@43875
   523
  in
blanchet@46250
   524
    if timeout = Time.zeroTime then
blanchet@46391
   525
      Trust_Playable (get_preferred reconstrs, NONE)
blanchet@46250
   526
    else
blanchet@46391
   527
      play [] reconstrs
blanchet@43874
   528
  end
blanchet@43874
   529
blanchet@41561
   530
blanchet@41502
   531
(* generic TPTP-based ATPs *)
blanchet@40242
   532
blanchet@43595
   533
(* Too general means, positive equality literal with a variable X as one
blanchet@43595
   534
   operand, when X does not occur properly in the other operand. This rules out
blanchet@43595
   535
   clearly inconsistent facts such as X = a | X = b, though it by no means
blanchet@43595
   536
   guarantees soundness. *)
blanchet@43595
   537
blanchet@43595
   538
(* Unwanted equalities are those between a (bound or schematic) variable that
blanchet@43595
   539
   does not properly occur in the second operand. *)
blanchet@43595
   540
val is_exhaustive_finite =
blanchet@43595
   541
  let
blanchet@43595
   542
    fun is_bad_equal (Var z) t =
blanchet@43595
   543
        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
blanchet@43595
   544
      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
blanchet@43595
   545
      | is_bad_equal _ _ = false
blanchet@43595
   546
    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
blanchet@43595
   547
    fun do_formula pos t =
blanchet@43595
   548
      case (pos, t) of
blanchet@43595
   549
        (_, @{const Trueprop} $ t1) => do_formula pos t1
blanchet@43595
   550
      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
blanchet@43595
   551
        do_formula pos t'
blanchet@43595
   552
      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
blanchet@43595
   553
        do_formula pos t'
blanchet@43595
   554
      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
blanchet@43595
   555
        do_formula pos t'
blanchet@43595
   556
      | (_, @{const "==>"} $ t1 $ t2) =>
blanchet@43595
   557
        do_formula (not pos) t1 andalso
blanchet@43595
   558
        (t2 = @{prop False} orelse do_formula pos t2)
blanchet@43595
   559
      | (_, @{const HOL.implies} $ t1 $ t2) =>
blanchet@43595
   560
        do_formula (not pos) t1 andalso
blanchet@43595
   561
        (t2 = @{const False} orelse do_formula pos t2)
blanchet@43595
   562
      | (_, @{const Not} $ t1) => do_formula (not pos) t1
blanchet@43595
   563
      | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
blanchet@43595
   564
      | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
blanchet@43595
   565
      | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
blanchet@43595
   566
      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
blanchet@43595
   567
      | _ => false
blanchet@43595
   568
  in do_formula true end
blanchet@43595
   569
blanchet@43595
   570
fun has_bound_or_var_of_type pred =
blanchet@43595
   571
  exists_subterm (fn Var (_, T as Type _) => pred T
blanchet@43595
   572
                   | Abs (_, T as Type _, _) => pred T
blanchet@43595
   573
                   | _ => false)
blanchet@43595
   574
blanchet@43595
   575
(* Facts are forbidden to contain variables of these types. The typical reason
blanchet@43595
   576
   is that they lead to unsoundness. Note that "unit" satisfies numerous
blanchet@43595
   577
   equations like "?x = ()". The resulting clauses will have no type constraint,
blanchet@43595
   578
   yielding false proofs. Even "bool" leads to many unsound proofs, though only
blanchet@43595
   579
   for higher-order problems. *)
blanchet@43595
   580
blanchet@43595
   581
(* Facts containing variables of type "unit" or "bool" or of the form
blanchet@43595
   582
   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
blanchet@43595
   583
   are omitted. *)
blanchet@43785
   584
fun is_dangerous_prop ctxt =
blanchet@43785
   585
  transform_elim_prop
blanchet@45252
   586
  #> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf
blanchet@43595
   587
      is_exhaustive_finite)
blanchet@43595
   588
blanchet@39732
   589
(* Important messages are important but not so important that users want to see
blanchet@39732
   590
   them each time. *)
blanchet@45513
   591
val atp_important_message_keep_quotient = 25
blanchet@39732
   592
blanchet@45275
   593
fun choose_type_enc soundness best_type_enc format =
blanchet@45256
   594
  the_default best_type_enc
blanchet@45256
   595
  #> type_enc_from_string soundness
blanchet@45275
   596
  #> adjust_type_enc format
blanchet@43419
   597
blanchet@43892
   598
val metis_minimize_max_time = seconds 2.0
blanchet@43892
   599
blanchet@46432
   600
fun choose_minimize_command params minimize_command name preplay =
blanchet@46391
   601
  let
blanchet@46391
   602
    val (name, override_params) =
blanchet@46391
   603
      case preplay of
blanchet@46391
   604
        Played (reconstr, time) =>
blanchet@46391
   605
        if Time.<= (time, metis_minimize_max_time) then
blanchet@46432
   606
          extract_reconstructor params reconstr
blanchet@46391
   607
        else
blanchet@46391
   608
          (name, [])
blanchet@46391
   609
      | _ => (name, [])
blanchet@46391
   610
  in minimize_command override_params name end
blanchet@43892
   611
blanchet@48977
   612
fun repair_monomorph_context max_iters best_max_iters max_new_instances
blanchet@48977
   613
                             best_max_new_instances =
blanchet@48977
   614
  Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters)
blanchet@48977
   615
  #> Config.put Monomorph.max_new_instances
blanchet@48977
   616
         (max_new_instances |> the_default best_max_new_instances)
boehmes@44071
   617
  #> Config.put Monomorph.keep_partial_instances false
blanchet@44067
   618
blanchet@45364
   619
fun suffix_for_mode Auto_Try = "_auto_try"
blanchet@45364
   620
  | suffix_for_mode Try = "_try"
blanchet@45364
   621
  | suffix_for_mode Normal = ""
blanchet@46445
   622
  | suffix_for_mode Auto_Minimize = "_auto_min"
blanchet@45364
   623
  | suffix_for_mode Minimize = "_min"
blanchet@45364
   624
blanchet@45282
   625
(* Give the ATPs some slack before interrupting them the hard way. "z3_tptp" on
blanchet@44498
   626
   Linux appears to be the only ATP that does not honor its time limit. *)
blanchet@44556
   627
val atp_timeout_slack = seconds 1.0
blanchet@44498
   628
blanchet@43862
   629
fun run_atp mode name
blanchet@47925
   630
        ({exec, required_vars, arguments, proof_delims, known_failures,
blanchet@48991
   631
          prem_role, best_slices, best_max_mono_iters,
blanchet@48977
   632
          best_max_new_mono_instances, ...} : atp_config)
blanchet@47129
   633
        (params as {debug, verbose, overlord, type_enc, strict, lam_trans,
blanchet@49308
   634
                    uncurried_aliases, max_facts, max_mono_iters,
blanchet@47237
   635
                    max_new_mono_instances, isar_proof, isar_shrink_factor,
blanchet@47237
   636
                    slice, timeout, preplay_timeout, ...})
blanchet@43878
   637
        minimize_command
blanchet@43878
   638
        ({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) =
blanchet@38257
   639
  let
blanchet@43053
   640
    val thy = Proof.theory_of state
blanchet@39564
   641
    val ctxt = Proof.context_of state
blanchet@48961
   642
    val atp_mode =
blanchet@49158
   643
      if Config.get ctxt completish then Sledgehammer_Completish
blanchet@48961
   644
      else Sledgehammer
blanchet@43845
   645
    val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal
blanchet@41407
   646
    val (dest_dir, problem_prefix) =
blanchet@41407
   647
      if overlord then overlord_file_location_for_prover name
blanchet@41407
   648
      else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
blanchet@40242
   649
    val problem_file_name =
blanchet@41407
   650
      Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
blanchet@45364
   651
                  suffix_for_mode mode ^ "_" ^ string_of_int subgoal)
blanchet@40242
   652
    val problem_path_name =
blanchet@40240
   653
      if dest_dir = "" then
blanchet@40242
   654
        File.tmp_path problem_file_name
blanchet@40240
   655
      else if File.exists (Path.explode dest_dir) then
blanchet@40242
   656
        Path.append (Path.explode dest_dir) problem_file_name
blanchet@39247
   657
      else
blanchet@40240
   658
        error ("No such directory: " ^ quote dest_dir ^ ".")
blanchet@47925
   659
    val command =
blanchet@47925
   660
      case find_first (fn var => getenv var <> "") (fst exec) of
blanchet@47925
   661
        SOME var => Path.explode (getenv var ^ "/" ^ snd exec)
blanchet@47925
   662
      | NONE => error ("The environment variable " ^ quote (hd (fst exec)) ^
blanchet@47925
   663
                       " is not set.")
blanchet@38257
   664
    fun split_time s =
blanchet@38257
   665
      let
blanchet@43319
   666
        val split = String.tokens (fn c => str c = "\n")
blanchet@48603
   667
        val (output, t) =
blanchet@48603
   668
          s |> split |> (try split_last #> the_default ([], "0"))
blanchet@48603
   669
            |>> cat_lines
blanchet@43319
   670
        fun as_num f = f >> (fst o read_int)
blanchet@43319
   671
        val num = as_num (Scan.many1 Symbol.is_ascii_digit)
blanchet@43319
   672
        val digit = Scan.one Symbol.is_ascii_digit
blanchet@43319
   673
        val num3 = as_num (digit ::: digit ::: (digit >> single))
blanchet@43319
   674
        val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
blanchet@46252
   675
        val as_time =
blanchet@46252
   676
          raw_explode #> Scan.read Symbol.stopper time #> the_default 0
blanchet@48603
   677
      in (output, as_time t |> Time.fromMilliseconds) end
blanchet@41561
   678
    fun run_on prob_file =
blanchet@47925
   679
      case find_first (forall (fn var => getenv var = ""))
blanchet@47925
   680
                      (fst exec :: required_vars) of
blanchet@47925
   681
        SOME home_vars =>
blanchet@47925
   682
        error ("The environment variable " ^ quote (hd home_vars) ^
blanchet@47925
   683
               " is not set.")
blanchet@47925
   684
      | NONE =>
blanchet@38278
   685
        if File.exists command then
blanchet@38278
   686
          let
blanchet@43314
   687
            (* If slicing is disabled, we expand the last slice to fill the
blanchet@43314
   688
               entire time available. *)
blanchet@46577
   689
            val actual_slices = get_slices slice (best_slices ctxt)
blanchet@43314
   690
            val num_actual_slices = length actual_slices
blanchet@43316
   691
            fun monomorphize_facts facts =
blanchet@43316
   692
              let
blanchet@44067
   693
                val ctxt =
blanchet@48977
   694
                  ctxt
blanchet@48977
   695
                  |> repair_monomorph_context max_mono_iters best_max_mono_iters
blanchet@48977
   696
                          max_new_mono_instances best_max_new_mono_instances
blanchet@43316
   697
                (* pseudo-theorem involving the same constants as the subgoal *)
blanchet@43316
   698
                val subgoal_th =
blanchet@43316
   699
                  Logic.list_implies (hyp_ts, concl_t)
blanchet@43316
   700
                  |> Skip_Proof.make_thm thy
blanchet@44090
   701
                val rths =
blanchet@44090
   702
                  facts |> chop (length facts div 4)
blanchet@44090
   703
                        |>> map (pair 1 o snd)
blanchet@44090
   704
                        ||> map (pair 2 o snd)
blanchet@44090
   705
                        |> op @
blanchet@44090
   706
                        |> cons (0, subgoal_th)
blanchet@43316
   707
              in
blanchet@44090
   708
                Monomorph.monomorph atp_schematic_consts_of rths ctxt
blanchet@44067
   709
                |> fst |> tl
blanchet@44067
   710
                |> curry ListPair.zip (map fst facts)
blanchet@47909
   711
                |> maps (fn (name, rths) =>
blanchet@47909
   712
                            map (pair name o zero_var_indexes o snd) rths)
blanchet@43316
   713
              end
blanchet@47235
   714
            fun run_slice time_left (cache_key, cache_value)
blanchet@47235
   715
                    (slice, (time_frac, (complete,
blanchet@49308
   716
                        (key as (best_max_facts, format, best_type_enc,
blanchet@47237
   717
                                 best_lam_trans, best_uncurried_aliases),
blanchet@47237
   718
                                 extra)))) =
blanchet@38278
   719
              let
blanchet@43314
   720
                val num_facts =
blanchet@49308
   721
                  length facts |> is_none max_facts ? Integer.min best_max_facts
blanchet@47129
   722
                val soundness = if strict then Strict else Non_Strict
blanchet@45275
   723
                val type_enc =
blanchet@45275
   724
                  type_enc |> choose_type_enc soundness best_type_enc format
blanchet@49104
   725
                val sound = is_type_enc_sound type_enc
blanchet@43314
   726
                val real_ms = Real.fromInt o Time.toMilliseconds
blanchet@43314
   727
                val slice_timeout =
blanchet@43314
   728
                  ((real_ms time_left
blanchet@43314
   729
                    |> (if slice < num_actual_slices - 1 then
blanchet@43314
   730
                          curry Real.min (time_frac * real_ms timeout)
blanchet@43314
   731
                        else
blanchet@43314
   732
                          I))
blanchet@43314
   733
                   * 0.001) |> seconds
blanchet@44513
   734
                val generous_slice_timeout =
blanchet@44513
   735
                  Time.+ (slice_timeout, atp_timeout_slack)
blanchet@43314
   736
                val _ =
blanchet@43485
   737
                  if debug then
blanchet@43569
   738
                    quote name ^ " slice #" ^ string_of_int (slice + 1) ^
blanchet@43569
   739
                    " with " ^ string_of_int num_facts ^ " fact" ^
blanchet@43569
   740
                    plural_s num_facts ^ " for " ^
blanchet@43569
   741
                    string_from_time slice_timeout ^ "..."
blanchet@43314
   742
                    |> Output.urgent_message
blanchet@43314
   743
                  else
blanchet@43314
   744
                    ()
blanchet@46385
   745
                val readable_names = not (Config.get ctxt atp_full_names)
blanchet@46391
   746
                val lam_trans =
blanchet@46391
   747
                  case lam_trans of
blanchet@46391
   748
                    SOME s => s
blanchet@46392
   749
                  | NONE => best_lam_trans
blanchet@47237
   750
                val uncurried_aliases =
blanchet@47237
   751
                  case uncurried_aliases of
blanchet@47237
   752
                    SOME b => b
blanchet@47237
   753
                  | NONE => best_uncurried_aliases
blanchet@47235
   754
                val value as (atp_problem, _, fact_names, _, _) =
blanchet@47235
   755
                  if cache_key = SOME key then
blanchet@47235
   756
                    cache_value
blanchet@47235
   757
                  else
blanchet@47235
   758
                    facts
blanchet@47235
   759
                    |> map untranslated_fact
blanchet@49104
   760
                    |> not sound
blanchet@47235
   761
                       ? filter_out (is_dangerous_prop ctxt o prop_of o snd)
blanchet@47235
   762
                    |> take num_facts
blanchet@49146
   763
                    |> not (is_type_enc_polymorphic type_enc)
blanchet@47235
   764
                       ? monomorphize_facts
blanchet@47235
   765
                    |> map (apsnd prop_of)
blanchet@48991
   766
                    |> prepare_atp_problem ctxt format prem_role type_enc
blanchet@48961
   767
                                           atp_mode lam_trans uncurried_aliases
blanchet@48961
   768
                                           readable_names true hyp_ts concl_t
blanchet@47901
   769
                fun sel_weights () = atp_problem_selection_weights atp_problem
blanchet@47909
   770
                fun ord_info () = atp_problem_term_order_info atp_problem
blanchet@47909
   771
                val ord = effective_term_order ctxt name
blanchet@44231
   772
                val full_proof = debug orelse isar_proof
blanchet@47909
   773
                val args = arguments ctxt full_proof extra slice_timeout
blanchet@47909
   774
                                     (ord, ord_info, sel_weights)
blanchet@47235
   775
                val command =
blanchet@47909
   776
                  File.shell_path command ^ " " ^ args ^ " " ^
blanchet@47909
   777
                  File.shell_path prob_file
blanchet@47235
   778
                  |> enclose "TIMEFORMAT='%3R'; { time " " ; } 2>&1"
blanchet@47235
   779
                val _ =
blanchet@47270
   780
                  atp_problem
blanchet@47909
   781
                  |> lines_for_atp_problem format ord ord_info
blanchet@47270
   782
                  |> cons ("% " ^ command ^ "\n")
blanchet@47270
   783
                  |> File.write_list prob_file
blanchet@46252
   784
                val ((output, run_time), (atp_proof, outcome)) =
blanchet@45494
   785
                  TimeLimit.timeLimit generous_slice_timeout
blanchet@45494
   786
                                      Isabelle_System.bash_output command
blanchet@38278
   787
                  |>> (if overlord then
blanchet@38278
   788
                         prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
blanchet@38278
   789
                       else
blanchet@38278
   790
                         I)
blanchet@45494
   791
                  |> fst |> split_time
blanchet@44498
   792
                  |> (fn accum as (output, _) =>
blanchet@44498
   793
                         (accum,
blanchet@44498
   794
                          extract_tstplike_proof_and_outcome verbose complete
blanchet@44498
   795
                              proof_delims known_failures output
blanchet@49145
   796
                          |>> atp_proof_from_tstplike_proof atp_problem
blanchet@44498
   797
                          handle UNRECOGNIZED_ATP_PROOF () =>
blanchet@44498
   798
                                 ([], SOME ProofIncomplete)))
blanchet@44498
   799
                  handle TimeLimit.TimeOut =>
blanchet@46252
   800
                         (("", slice_timeout), ([], SOME TimedOut))
blanchet@43320
   801
                val outcome =
blanchet@43322
   802
                  case outcome of
blanchet@43458
   803
                    NONE =>
blanchet@46422
   804
                    (case used_facts_in_unsound_atp_proof ctxt fact_names
blanchet@46422
   805
                                                          atp_proof
blanchet@45729
   806
                          |> Option.map (sort string_ord) of
blanchet@45729
   807
                       SOME facts =>
blanchet@45729
   808
                       let
blanchet@45729
   809
                         val failure =
blanchet@49119
   810
                           UnsoundProof (is_type_enc_sound type_enc, facts)
blanchet@45729
   811
                       in
blanchet@45729
   812
                         if debug then
blanchet@45729
   813
                           (warning (string_for_failure failure); NONE)
blanchet@45729
   814
                         else
blanchet@45729
   815
                           SOME failure
blanchet@45729
   816
                       end
blanchet@45729
   817
                     | NONE => NONE)
blanchet@43322
   818
                  | _ => outcome
blanchet@47235
   819
              in ((SOME key, value), (output, run_time, atp_proof, outcome)) end
blanchet@38883
   820
            val timer = Timer.startRealTimer ()
blanchet@46252
   821
            fun maybe_run_slice slice
blanchet@47235
   822
                    (result as (cache, (_, run_time0, _, SOME _))) =
blanchet@43323
   823
                let
blanchet@43323
   824
                  val time_left = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@43323
   825
                in
blanchet@43323
   826
                  if Time.<= (time_left, Time.zeroTime) then
blanchet@43323
   827
                    result
blanchet@43323
   828
                  else
blanchet@47235
   829
                    run_slice time_left cache slice
blanchet@47235
   830
                    |> (fn (cache, (output, run_time, atp_proof, outcome)) =>
blanchet@47235
   831
                           (cache, (output, Time.+ (run_time0, run_time),
blanchet@46252
   832
                                    atp_proof, outcome)))
blanchet@43323
   833
                end
blanchet@44351
   834
              | maybe_run_slice _ result = result
blanchet@43314
   835
          in
blanchet@47235
   836
            ((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)),
blanchet@46422
   837
             ("", Time.zeroTime, [], SOME InternalError))
blanchet@44351
   838
            |> fold maybe_run_slice actual_slices
blanchet@43314
   839
          end
blanchet@38278
   840
        else
wenzelm@44475
   841
          error ("Bad executable: " ^ Path.print command)
blanchet@38257
   842
blanchet@38257
   843
    (* If the problem file has not been exported, remove it; otherwise, export
blanchet@38257
   844
       the proof file too. *)
blanchet@41561
   845
    fun cleanup prob_file =
blanchet@41561
   846
      if dest_dir = "" then try File.rm prob_file else NONE
blanchet@44175
   847
    fun export prob_file (_, (output, _, _, _)) =
blanchet@40240
   848
      if dest_dir = "" then
blanchet@38257
   849
        ()
blanchet@38257
   850
      else
blanchet@41561
   851
        File.write (Path.explode (Path.implode prob_file ^ "_proof")) output
blanchet@47235
   852
    val ((_, (_, pool, fact_names, _, sym_tab)),
blanchet@47235
   853
         (output, run_time, atp_proof, outcome)) =
blanchet@40242
   854
      with_path cleanup export run_on problem_path_name
blanchet@39732
   855
    val important_message =
blanchet@43862
   856
      if mode = Normal andalso
blanchet@43480
   857
         random_range 0 (atp_important_message_keep_quotient - 1) = 0 then
blanchet@39732
   858
        extract_important_message output
blanchet@39732
   859
      else
blanchet@39732
   860
        ""
blanchet@44102
   861
    val (used_facts, preplay, message, message_tail) =
blanchet@38257
   862
      case outcome of
blanchet@38257
   863
        NONE =>
blanchet@43874
   864
        let
blanchet@46422
   865
          val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof
blanchet@46454
   866
          val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof
blanchet@46392
   867
          val reconstrs =
blanchet@46425
   868
            bunch_of_reconstructors needs_full_types
blanchet@46431
   869
                (lam_trans_from_atp_proof atp_proof
blanchet@47233
   870
                 o (fn desperate => if desperate then hide_lamsN
blanchet@47233
   871
                                    else metis_default_lam_trans))
blanchet@43874
   872
        in
blanchet@43893
   873
          (used_facts,
blanchet@43893
   874
           fn () =>
blanchet@43893
   875
              let
blanchet@46652
   876
                val used_pairs =
blanchet@43893
   877
                  facts |> map untranslated_fact |> filter_used_facts used_facts
blanchet@43893
   878
              in
blanchet@46652
   879
                play_one_line_proof mode debug verbose preplay_timeout
blanchet@46652
   880
                    used_pairs state subgoal (hd reconstrs) reconstrs
blanchet@43893
   881
              end,
blanchet@43893
   882
           fn preplay =>
blanchet@43893
   883
              let
blanchet@43893
   884
                val isar_params =
blanchet@46423
   885
                  (debug, isar_shrink_factor, pool, fact_names, sym_tab,
blanchet@46423
   886
                   atp_proof, goal)
blanchet@43893
   887
                val one_line_params =
blanchet@43893
   888
                  (preplay, proof_banner mode name, used_facts,
blanchet@46432
   889
                   choose_minimize_command params minimize_command name preplay,
blanchet@43893
   890
                   subgoal, subgoal_count)
blanchet@44102
   891
              in proof_text ctxt isar_proof isar_params one_line_params end,
blanchet@44102
   892
           (if verbose then
blanchet@46252
   893
              "\nATP real CPU time: " ^ string_from_time run_time ^ "."
blanchet@44102
   894
            else
blanchet@44102
   895
              "") ^
blanchet@44102
   896
           (if important_message <> "" then
blanchet@44102
   897
              "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^
blanchet@44102
   898
              important_message
blanchet@44102
   899
            else
blanchet@44102
   900
              ""))
blanchet@43874
   901
        end
blanchet@43893
   902
      | SOME failure =>
blanchet@46390
   903
        ([], K (Failed_to_Play plain_metis),
blanchet@46390
   904
         fn _ => string_for_failure failure, "")
blanchet@38257
   905
  in
blanchet@46252
   906
    {outcome = outcome, used_facts = used_facts, run_time = run_time,
blanchet@44102
   907
     preplay = preplay, message = message, message_tail = message_tail}
blanchet@38257
   908
  end
blanchet@38257
   909
blanchet@40917
   910
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
blanchet@40917
   911
   these are sorted out properly in the SMT module, we have to interpret these
blanchet@40917
   912
   ourselves. *)
blanchet@40932
   913
val remote_smt_failures =
blanchet@44498
   914
  [(2, NoLibwwwPerl),
blanchet@44498
   915
   (22, CantConnect)]
blanchet@40932
   916
val z3_failures =
blanchet@41477
   917
  [(101, OutOfResources),
blanchet@41477
   918
   (103, MalformedInput),
blanchet@41470
   919
   (110, MalformedInput)]
blanchet@40932
   920
val unix_failures =
blanchet@40932
   921
  [(139, Crashed)]
blanchet@44498
   922
val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
blanchet@40797
   923
blanchet@42964
   924
fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
blanchet@43891
   925
    if is_real_cex then Unprovable else GaveUp
blanchet@41470
   926
  | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
blanchet@41470
   927
  | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
blanchet@41470
   928
    (case AList.lookup (op =) smt_failures code of
blanchet@40932
   929
       SOME failure => failure
blanchet@41505
   930
     | NONE => UnknownError ("Abnormal termination with exit code " ^
blanchet@41505
   931
                             string_of_int code ^ "."))
blanchet@41470
   932
  | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
blanchet@41470
   933
  | failure_from_smt_failure (SMT_Failure.Other_Failure msg) =
blanchet@42916
   934
    UnknownError msg
wenzelm@28586
   935
blanchet@40946
   936
(* FUDGE *)
blanchet@43517
   937
val smt_max_slices =
blanchet@43517
   938
  Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
blanchet@43517
   939
val smt_slice_fact_frac =
blanchet@43517
   940
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
blanchet@43517
   941
val smt_slice_time_frac =
blanchet@43517
   942
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
blanchet@43517
   943
val smt_slice_min_secs =
blanchet@43517
   944
  Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
blanchet@40636
   945
blanchet@43517
   946
fun smt_filter_loop ctxt name
blanchet@43589
   947
                    ({debug, verbose, overlord, max_mono_iters,
blanchet@46577
   948
                      max_new_mono_instances, timeout, slice, ...} : params)
blanchet@48402
   949
                    state i =
blanchet@40636
   950
  let
blanchet@46577
   951
    val max_slices = if slice then Config.get ctxt smt_max_slices else 1
blanchet@41483
   952
    val repair_context =
blanchet@44074
   953
      select_smt_solver name
blanchet@46428
   954
      #> Config.put SMT_Config.verbose debug
blanchet@44074
   955
      #> (if overlord then
blanchet@44074
   956
            Config.put SMT_Config.debug_files
blanchet@44074
   957
                       (overlord_file_location_for_prover name
blanchet@44074
   958
                        |> (fn (path, name) => path ^ "/" ^ name))
blanchet@44074
   959
          else
blanchet@44074
   960
            I)
blanchet@44074
   961
      #> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers)
blanchet@41483
   962
    val state = state |> Proof.map_context repair_context
blanchet@43314
   963
    fun do_slice timeout slice outcome0 time_so_far facts =
blanchet@40795
   964
      let
blanchet@40795
   965
        val timer = Timer.startRealTimer ()
blanchet@43589
   966
        val state =
blanchet@43589
   967
          state |> Proof.map_context
blanchet@44108
   968
                       (repair_monomorph_context max_mono_iters
blanchet@48977
   969
                            default_max_mono_iters max_new_mono_instances
blanchet@48977
   970
                            default_max_new_mono_instances)
blanchet@40795
   971
        val ms = timeout |> Time.toMilliseconds
blanchet@43314
   972
        val slice_timeout =
blanchet@43314
   973
          if slice < max_slices then
blanchet@41415
   974
            Int.min (ms,
blanchet@43517
   975
                Int.max (1000 * Config.get ctxt smt_slice_min_secs,
blanchet@43517
   976
                    Real.ceil (Config.get ctxt smt_slice_time_frac
blanchet@43517
   977
                               * Real.fromInt ms)))
blanchet@40795
   978
            |> Time.fromMilliseconds
blanchet@40795
   979
          else
blanchet@40795
   980
            timeout
blanchet@40795
   981
        val num_facts = length facts
blanchet@40795
   982
        val _ =
blanchet@43485
   983
          if debug then
blanchet@43485
   984
            quote name ^ " slice " ^ string_of_int slice ^ " with " ^
blanchet@43485
   985
            string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^
blanchet@43485
   986
            string_from_time slice_timeout ^ "..."
blanchet@40795
   987
            |> Output.urgent_message
blanchet@40795
   988
          else
blanchet@40795
   989
            ()
blanchet@41414
   990
        val birth = Timer.checkRealTimer timer
blanchet@41419
   991
        val _ =
blanchet@41459
   992
          if debug then Output.urgent_message "Invoking SMT solver..." else ()
blanchet@41457
   993
        val (outcome, used_facts) =
blanchet@48402
   994
          SMT_Solver.smt_filter_preprocess state facts i
blanchet@43314
   995
          |> SMT_Solver.smt_filter_apply slice_timeout
blanchet@41480
   996
          |> (fn {outcome, used_facts} => (outcome, used_facts))
blanchet@41457
   997
          handle exn => if Exn.is_interrupt exn then
blanchet@41457
   998
                          reraise exn
blanchet@41457
   999
                        else
blanchet@42916
  1000
                          (ML_Compiler.exn_message exn
blanchet@41457
  1001
                           |> SMT_Failure.Other_Failure |> SOME, [])
blanchet@41414
  1002
        val death = Timer.checkRealTimer timer
blanchet@40795
  1003
        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
blanchet@41414
  1004
        val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
blanchet@40795
  1005
        val too_many_facts_perhaps =
blanchet@40795
  1006
          case outcome of
blanchet@40795
  1007
            NONE => false
blanchet@40795
  1008
          | SOME (SMT_Failure.Counterexample _) => false
blanchet@43314
  1009
          | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
blanchet@43485
  1010
          | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
blanchet@40795
  1011
          | SOME SMT_Failure.Out_Of_Memory => true
blanchet@41459
  1012
          | SOME (SMT_Failure.Other_Failure _) => true
blanchet@40795
  1013
        val timeout = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@40795
  1014
      in
blanchet@43314
  1015
        if too_many_facts_perhaps andalso slice < max_slices andalso
blanchet@40795
  1016
           num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
blanchet@41415
  1017
          let
blanchet@43485
  1018
            val new_num_facts =
blanchet@43517
  1019
              Real.ceil (Config.get ctxt smt_slice_fact_frac
blanchet@43517
  1020
                         * Real.fromInt num_facts)
blanchet@43485
  1021
            val _ =
blanchet@43485
  1022
              if verbose andalso is_some outcome then
blanchet@43485
  1023
                quote name ^ " invoked with " ^ string_of_int num_facts ^
blanchet@43485
  1024
                " fact" ^ plural_s num_facts ^ ": " ^
blanchet@43485
  1025
                string_for_failure (failure_from_smt_failure (the outcome)) ^
blanchet@43485
  1026
                " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^
blanchet@43509
  1027
                plural_s new_num_facts ^ "..."
blanchet@43485
  1028
                |> Output.urgent_message
blanchet@43485
  1029
              else
blanchet@43485
  1030
                ()
blanchet@43314
  1031
          in
blanchet@43485
  1032
            facts |> take new_num_facts
blanchet@43485
  1033
                  |> do_slice timeout (slice + 1) outcome0 time_so_far
blanchet@43314
  1034
          end
blanchet@40795
  1035
        else
blanchet@40795
  1036
          {outcome = if is_none outcome then NONE else the outcome0,
blanchet@46231
  1037
           used_facts = used_facts, run_time = time_so_far}
blanchet@40636
  1038
      end
blanchet@43314
  1039
  in do_slice timeout 1 NONE Time.zeroTime end
blanchet@40636
  1040
blanchet@43893
  1041
fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
blanchet@43852
  1042
        minimize_command
blanchet@48402
  1043
        ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
blanchet@36379
  1044
  let
blanchet@41483
  1045
    val ctxt = Proof.context_of state
blanchet@41502
  1046
    val num_facts = length facts
blanchet@41502
  1047
    val facts = facts ~~ (0 upto num_facts - 1)
blanchet@43517
  1048
                |> map (smt_weighted_fact ctxt num_facts)
blanchet@46652
  1049
    val {outcome, used_facts = used_pairs, run_time} =
blanchet@48402
  1050
      smt_filter_loop ctxt name params state subgoal facts
blanchet@46652
  1051
    val used_facts = used_pairs |> map fst
blanchet@41470
  1052
    val outcome = outcome |> Option.map failure_from_smt_failure
blanchet@44102
  1053
    val (preplay, message, message_tail) =
blanchet@40425
  1054
      case outcome of
blanchet@40425
  1055
        NONE =>
blanchet@43893
  1056
        (fn () =>
blanchet@46652
  1057
            play_one_line_proof mode debug verbose preplay_timeout used_pairs
blanchet@46425
  1058
                state subgoal SMT
blanchet@46431
  1059
                (bunch_of_reconstructors false
blanchet@46431
  1060
                     (fn plain =>
blanchet@47193
  1061
                         if plain then metis_default_lam_trans else liftingN)),
blanchet@43893
  1062
         fn preplay =>
blanchet@43893
  1063
            let
blanchet@43893
  1064
              val one_line_params =
blanchet@43893
  1065
                (preplay, proof_banner mode name, used_facts,
blanchet@46432
  1066
                 choose_minimize_command params minimize_command name preplay,
blanchet@43893
  1067
                 subgoal, subgoal_count)
blanchet@44102
  1068
            in one_line_proof_text one_line_params end,
blanchet@44102
  1069
         if verbose then
blanchet@46231
  1070
           "\nSMT solver real CPU time: " ^ string_from_time run_time ^ "."
blanchet@44102
  1071
         else
blanchet@44102
  1072
           "")
blanchet@44007
  1073
      | SOME failure =>
blanchet@46390
  1074
        (K (Failed_to_Play plain_metis), fn _ => string_for_failure failure, "")
blanchet@40244
  1075
  in
blanchet@46231
  1076
    {outcome = outcome, used_facts = used_facts, run_time = run_time,
blanchet@46231
  1077
     preplay = preplay, message = message, message_tail = message_tail}
blanchet@40244
  1078
  end
blanchet@40244
  1079
blanchet@46391
  1080
fun run_reconstructor mode name
blanchet@46432
  1081
        (params as {debug, verbose, timeout, type_enc, lam_trans, ...})
blanchet@46250
  1082
        minimize_command
blanchet@46250
  1083
        ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
blanchet@43891
  1084
  let
blanchet@46391
  1085
    val reconstr =
blanchet@46391
  1086
      if name = metisN then
blanchet@46391
  1087
        Metis (type_enc |> the_default (hd partial_type_encs),
blanchet@46391
  1088
               lam_trans |> the_default metis_default_lam_trans)
blanchet@46391
  1089
      else if name = smtN then
blanchet@46391
  1090
        SMT
blanchet@46391
  1091
      else
blanchet@46391
  1092
        raise Fail ("unknown reconstructor: " ^ quote name)
blanchet@46652
  1093
    val used_pairs = facts |> map untranslated_fact
blanchet@46652
  1094
    val used_facts = used_pairs |> map fst
blanchet@43891
  1095
  in
blanchet@46250
  1096
    case play_one_line_proof (if mode = Minimize then Normal else mode) debug
blanchet@46652
  1097
                             verbose timeout used_pairs state subgoal reconstr
blanchet@46391
  1098
                             [reconstr] of
blanchet@43891
  1099
      play as Played (_, time) =>
blanchet@46231
  1100
      {outcome = NONE, used_facts = used_facts, run_time = time,
blanchet@43893
  1101
       preplay = K play,
blanchet@46432
  1102
       message =
blanchet@46432
  1103
         fn play =>
blanchet@46432
  1104
            let
blanchet@46432
  1105
              val (_, override_params) = extract_reconstructor params reconstr
blanchet@46432
  1106
              val one_line_params =
blanchet@46432
  1107
                (play, proof_banner mode name, used_facts,
blanchet@46432
  1108
                 minimize_command override_params name, subgoal,
blanchet@46432
  1109
                 subgoal_count)
blanchet@46432
  1110
            in one_line_proof_text one_line_params end,
blanchet@44102
  1111
       message_tail = ""}
blanchet@43893
  1112
    | play =>
blanchet@44007
  1113
      let
blanchet@44007
  1114
        val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut
blanchet@44007
  1115
      in
blanchet@46231
  1116
        {outcome = SOME failure, used_facts = [], run_time = Time.zeroTime,
blanchet@44102
  1117
         preplay = K play, message = fn _ => string_for_failure failure,
blanchet@44102
  1118
         message_tail = ""}
blanchet@43891
  1119
      end
blanchet@43891
  1120
  end
blanchet@43891
  1121
blanchet@43862
  1122
fun get_prover ctxt mode name =
wenzelm@43232
  1123
  let val thy = Proof_Context.theory_of ctxt in
blanchet@46250
  1124
    if is_reconstructor name then run_reconstructor mode name
blanchet@48469
  1125
    else if is_atp thy name then run_atp mode name (get_atp thy name ())
blanchet@43893
  1126
    else if is_smt_prover ctxt name then run_smt_solver mode name
blanchet@43893
  1127
    else error ("No such prover: " ^ name ^ ".")
blanchet@41189
  1128
  end
blanchet@40244
  1129
wenzelm@28582
  1130
end;