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