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