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