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