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