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