src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML
author blanchet
Thu, 19 May 2011 10:24:13 +0200
changeset 43718 ba45312308b5
parent 43708 358769224d94
child 43722 de1fe9eaf3f4
permissions -rw-r--r--
reintroduce TFF workaround for limitations of actual TFF implementations (ToFoF, SNARK)
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@39232
    12
  type locality = Sledgehammer_Filter.locality
blanchet@40251
    13
  type relevance_fudge = Sledgehammer_Filter.relevance_fudge
blanchet@40358
    14
  type translated_formula = Sledgehammer_ATP_Translate.translated_formula
blanchet@41386
    15
  type type_system = Sledgehammer_ATP_Translate.type_system
blanchet@40249
    16
  type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
blanchet@39733
    17
blanchet@43450
    18
  datatype rich_type_system =
blanchet@43450
    19
    Dumb_Type_Sys of type_system |
blanchet@43450
    20
    Smart_Type_Sys of bool
blanchet@43450
    21
blanchet@35967
    22
  type params =
blanchet@41456
    23
    {debug: bool,
blanchet@35967
    24
     verbose: bool,
blanchet@36141
    25
     overlord: bool,
blanchet@41456
    26
     blocking: bool,
blanchet@40240
    27
     provers: string list,
blanchet@43450
    28
     type_sys: rich_type_system,
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@36235
    33
     explicit_apply: bool,
blanchet@35967
    34
     isar_proof: bool,
blanchet@36916
    35
     isar_shrink_factor: int,
blanchet@43314
    36
     slicing: bool,
blanchet@39229
    37
     timeout: Time.time,
blanchet@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@40243
    56
     message: string}
blanchet@39733
    57
blanchet@40242
    58
  type prover = params -> minimize_command -> prover_problem -> prover_result
blanchet@35867
    59
blanchet@43517
    60
  val smt_triggers : bool Config.T
blanchet@43517
    61
  val smt_weights : bool Config.T
blanchet@43517
    62
  val smt_weight_min_facts : int Config.T
blanchet@43517
    63
  val smt_min_weight : int Config.T
blanchet@43517
    64
  val smt_max_weight : int Config.T
blanchet@43517
    65
  val smt_max_weight_index : int Config.T
blanchet@41502
    66
  val smt_weight_curve : (int -> int) Unsynchronized.ref
blanchet@43517
    67
  val smt_max_slices : int Config.T
blanchet@43517
    68
  val smt_slice_fact_frac : real Config.T
blanchet@43517
    69
  val smt_slice_time_frac : real Config.T
blanchet@43517
    70
  val smt_slice_min_secs : int Config.T
blanchet@41337
    71
  val das_Tool : string
blanchet@41483
    72
  val select_smt_solver : string -> Proof.context -> Proof.context
blanchet@41335
    73
  val is_smt_prover : Proof.context -> string -> bool
blanchet@42591
    74
  val is_prover_supported : Proof.context -> string -> bool
blanchet@40253
    75
  val is_prover_installed : Proof.context -> string -> bool
blanchet@43314
    76
  val default_max_relevant_for_prover : Proof.context -> bool -> string -> int
blanchet@40615
    77
  val is_built_in_const_for_prover :
blanchet@41584
    78
    Proof.context -> string -> string * typ -> term list -> bool * term list
blanchet@41335
    79
  val atp_relevance_fudge : relevance_fudge
blanchet@41335
    80
  val smt_relevance_fudge : relevance_fudge
blanchet@41189
    81
  val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
blanchet@38257
    82
  val dest_dir : string Config.T
blanchet@38257
    83
  val problem_prefix : string Config.T
blanchet@39247
    84
  val measure_run_time : bool Config.T
blanchet@41502
    85
  val weight_smt_fact :
blanchet@43517
    86
    Proof.context -> int -> ((string * locality) * thm) * int
blanchet@41502
    87
    -> (string * locality) * (int option * thm)
blanchet@41339
    88
  val untranslated_fact : prover_fact -> (string * locality) * thm
blanchet@41483
    89
  val smt_weighted_fact :
blanchet@43517
    90
    Proof.context -> int -> prover_fact * int
blanchet@41502
    91
    -> (string * locality) * (int option * thm)
blanchet@42591
    92
  val supported_provers : Proof.context -> unit
blanchet@40240
    93
  val kill_provers : unit -> unit
blanchet@40240
    94
  val running_provers : unit -> unit
blanchet@40240
    95
  val messages : int option -> unit
blanchet@43315
    96
  val get_prover : Proof.context -> bool -> string -> prover
wenzelm@28477
    97
end;
wenzelm@28477
    98
blanchet@41335
    99
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
wenzelm@28477
   100
struct
wenzelm@28477
   101
blanchet@38262
   102
open ATP_Problem
blanchet@39731
   103
open ATP_Proof
blanchet@38262
   104
open ATP_Systems
blanchet@39734
   105
open Metis_Translate
blanchet@38257
   106
open Sledgehammer_Util
blanchet@39232
   107
open Sledgehammer_Filter
blanchet@40249
   108
open Sledgehammer_ATP_Translate
blanchet@40249
   109
open Sledgehammer_ATP_Reconstruct
blanchet@37583
   110
blanchet@37583
   111
(** The Sledgehammer **)
blanchet@37583
   112
blanchet@38348
   113
(* Identifier to distinguish Sledgehammer from other tools using
blanchet@38348
   114
   "Async_Manager". *)
blanchet@37585
   115
val das_Tool = "Sledgehammer"
blanchet@37585
   116
blanchet@41483
   117
val select_smt_solver =
boehmes@41680
   118
  Context.proof_map o SMT_Config.select_solver
blanchet@41483
   119
blanchet@41189
   120
fun is_smt_prover ctxt name =
boehmes@41680
   121
  member (op =) (SMT_Solver.available_solvers_of ctxt) name
blanchet@40243
   122
blanchet@42591
   123
fun is_prover_supported ctxt name =
wenzelm@43232
   124
  let val thy = Proof_Context.theory_of ctxt in
blanchet@42591
   125
    is_smt_prover ctxt name orelse member (op =) (supported_atps thy) name
blanchet@41189
   126
  end
blanchet@40253
   127
boehmes@41680
   128
fun is_prover_installed ctxt =
wenzelm@43232
   129
  is_smt_prover ctxt orf is_atp_installed (Proof_Context.theory_of ctxt)
blanchet@41189
   130
blanchet@43314
   131
fun get_slices slicing slices =
blanchet@43318
   132
  (0 upto length slices - 1) ~~ slices |> not slicing ? (List.last #> single)
blanchet@43314
   133
blanchet@43314
   134
fun default_max_relevant_for_prover ctxt slicing name =
wenzelm@43232
   135
  let val thy = Proof_Context.theory_of ctxt in
blanchet@41228
   136
    if is_smt_prover ctxt name then
boehmes@41680
   137
      SMT_Solver.default_max_relevant ctxt name
blanchet@41228
   138
    else
blanchet@43588
   139
      fold (Integer.max o fst o snd o snd o snd)
blanchet@43517
   140
           (get_slices slicing (#best_slices (get_atp thy name) ctxt)) 0
blanchet@41189
   141
  end
blanchet@40244
   142
blanchet@41388
   143
(* These are either simplified away by "Meson.presimplify" (most of the time) or
blanchet@41388
   144
   handled specially via "fFalse", "fTrue", ..., "fequal". *)
blanchet@40252
   145
val atp_irrelevant_consts =
blanchet@41388
   146
  [@{const_name False}, @{const_name True}, @{const_name Not},
blanchet@41388
   147
   @{const_name conj}, @{const_name disj}, @{const_name implies},
blanchet@41388
   148
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
blanchet@40447
   149
blanchet@41483
   150
fun is_built_in_const_for_prover ctxt name =
blanchet@41483
   151
  if is_smt_prover ctxt name then
blanchet@41584
   152
    let val ctxt = ctxt |> select_smt_solver name in
blanchet@41584
   153
      fn x => fn ts =>
blanchet@41584
   154
         if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
blanchet@41584
   155
           (true, [])
blanchet@41584
   156
         else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
blanchet@41584
   157
           (true, ts)
blanchet@41584
   158
         else
blanchet@41584
   159
           (false, ts)
blanchet@41584
   160
    end
blanchet@41483
   161
  else
blanchet@41584
   162
    fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts)
blanchet@40252
   163
blanchet@40251
   164
(* FUDGE *)
blanchet@40251
   165
val atp_relevance_fudge =
blanchet@43603
   166
  {local_const_multiplier = 1.5,
blanchet@41407
   167
   worse_irrel_freq = 100.0,
blanchet@40251
   168
   higher_order_irrel_weight = 1.05,
blanchet@40251
   169
   abs_rel_weight = 0.5,
blanchet@40251
   170
   abs_irrel_weight = 2.0,
blanchet@43611
   171
   skolem_irrel_weight = 0.25,
blanchet@40251
   172
   theory_const_rel_weight = 0.5,
blanchet@40251
   173
   theory_const_irrel_weight = 0.25,
blanchet@43600
   174
   chained_const_irrel_weight = 0.25,
blanchet@40251
   175
   intro_bonus = 0.15,
blanchet@40251
   176
   elim_bonus = 0.15,
blanchet@40251
   177
   simp_bonus = 0.15,
blanchet@40251
   178
   local_bonus = 0.55,
blanchet@40251
   179
   assum_bonus = 1.05,
blanchet@40251
   180
   chained_bonus = 1.5,
blanchet@40251
   181
   max_imperfect = 11.5,
blanchet@40251
   182
   max_imperfect_exp = 1.0,
blanchet@40251
   183
   threshold_divisor = 2.0,
blanchet@41341
   184
   ridiculous_threshold = 0.01}
blanchet@40251
   185
blanchet@40252
   186
(* FUDGE (FIXME) *)
blanchet@40251
   187
val smt_relevance_fudge =
blanchet@43603
   188
  {local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
blanchet@41407
   189
   worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
blanchet@40252
   190
   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
blanchet@40252
   191
   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
blanchet@40252
   192
   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
blanchet@40252
   193
   skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
blanchet@40252
   194
   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
blanchet@40252
   195
   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
blanchet@43600
   196
   chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge,
blanchet@40252
   197
   intro_bonus = #intro_bonus atp_relevance_fudge,
blanchet@40252
   198
   elim_bonus = #elim_bonus atp_relevance_fudge,
blanchet@40252
   199
   simp_bonus = #simp_bonus atp_relevance_fudge,
blanchet@40252
   200
   local_bonus = #local_bonus atp_relevance_fudge,
blanchet@40252
   201
   assum_bonus = #assum_bonus atp_relevance_fudge,
blanchet@40252
   202
   chained_bonus = #chained_bonus atp_relevance_fudge,
blanchet@40252
   203
   max_imperfect = #max_imperfect atp_relevance_fudge,
blanchet@40252
   204
   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
blanchet@40252
   205
   threshold_divisor = #threshold_divisor atp_relevance_fudge,
blanchet@40252
   206
   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
blanchet@40251
   207
blanchet@41189
   208
fun relevance_fudge_for_prover ctxt name =
blanchet@41189
   209
  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
blanchet@40251
   210
blanchet@42591
   211
fun supported_provers ctxt =
blanchet@40241
   212
  let
wenzelm@43232
   213
    val thy = Proof_Context.theory_of ctxt
blanchet@40241
   214
    val (remote_provers, local_provers) =
blanchet@42591
   215
      sort_strings (supported_atps thy) @
blanchet@42591
   216
      sort_strings (SMT_Solver.available_solvers_of ctxt)
blanchet@40241
   217
      |> List.partition (String.isPrefix remote_prefix)
blanchet@40241
   218
  in
blanchet@42591
   219
    Output.urgent_message ("Supported provers: " ^
blanchet@40446
   220
                           commas (local_provers @ remote_provers) ^ ".")
blanchet@40241
   221
  end
blanchet@35967
   222
blanchet@40240
   223
fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
blanchet@40240
   224
fun running_provers () = Async_Manager.running_threads das_Tool "provers"
blanchet@40240
   225
val messages = Async_Manager.thread_messages das_Tool "prover"
blanchet@40240
   226
blanchet@40240
   227
(** problems, results, ATPs, etc. **)
blanchet@35967
   228
blanchet@43450
   229
datatype rich_type_system =
blanchet@43450
   230
  Dumb_Type_Sys of type_system |
blanchet@43450
   231
  Smart_Type_Sys of bool
blanchet@43450
   232
blanchet@35967
   233
type params =
blanchet@41456
   234
  {debug: bool,
blanchet@35967
   235
   verbose: bool,
blanchet@36141
   236
   overlord: bool,
blanchet@41456
   237
   blocking: bool,
blanchet@40240
   238
   provers: string list,
blanchet@43450
   239
   type_sys: rich_type_system,
blanchet@43051
   240
   relevance_thresholds: real * real,
blanchet@43051
   241
   max_relevant: int option,
blanchet@43589
   242
   max_mono_iters: int,
blanchet@43605
   243
   max_new_mono_instances: int,
blanchet@36235
   244
   explicit_apply: bool,
blanchet@35967
   245
   isar_proof: bool,
blanchet@36916
   246
   isar_shrink_factor: int,
blanchet@43314
   247
   slicing: bool,
blanchet@39229
   248
   timeout: Time.time,
blanchet@39229
   249
   expect: string}
blanchet@35867
   250
blanchet@41338
   251
datatype prover_fact =
blanchet@41338
   252
  Untranslated_Fact of (string * locality) * thm |
blanchet@41483
   253
  SMT_Weighted_Fact of (string * locality) * (int option * thm)
blanchet@40242
   254
blanchet@40242
   255
type prover_problem =
blanchet@39564
   256
  {state: Proof.state,
blanchet@39242
   257
   goal: thm,
blanchet@39242
   258
   subgoal: int,
blanchet@40246
   259
   subgoal_count: int,
blanchet@41483
   260
   facts: prover_fact list,
blanchet@42612
   261
   smt_filter: (string * locality) SMT_Solver.smt_filter_data option}
blanchet@35867
   262
blanchet@40242
   263
type prover_result =
blanchet@36370
   264
  {outcome: failure option,
blanchet@35967
   265
   message: string,
blanchet@40445
   266
   used_facts: (string * locality) list,
blanchet@40243
   267
   run_time_in_msecs: int option}
blanchet@35867
   268
blanchet@40242
   269
type prover = params -> minimize_command -> prover_problem -> prover_result
blanchet@35867
   270
blanchet@38257
   271
(* configuration attributes *)
wenzelm@28484
   272
wenzelm@43487
   273
val dest_dir =
wenzelm@43487
   274
  Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
blanchet@39235
   275
  (* Empty string means create files in Isabelle's temporary files directory. *)
wenzelm@28484
   276
wenzelm@43487
   277
val problem_prefix =
wenzelm@43487
   278
  Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
wenzelm@28477
   279
wenzelm@43487
   280
val measure_run_time =
wenzelm@43487
   281
  Attrib.setup_config_bool @{binding sledgehammer_measure_run_time} (K false)
wenzelm@32995
   282
blanchet@38257
   283
fun with_path cleanup after f path =
blanchet@38257
   284
  Exn.capture f path
blanchet@38257
   285
  |> tap (fn _ => cleanup path)
blanchet@38257
   286
  |> Exn.release
blanchet@38257
   287
  |> tap (after path)
wenzelm@28484
   288
blanchet@40241
   289
fun proof_banner auto =
blanchet@40465
   290
  if auto then "Auto Sledgehammer found a proof" else "Try this command"
blanchet@40241
   291
blanchet@43517
   292
val smt_triggers =
blanchet@43517
   293
  Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
blanchet@43517
   294
val smt_weights =
blanchet@43517
   295
  Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
blanchet@43517
   296
val smt_weight_min_facts =
blanchet@43517
   297
  Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
blanchet@41502
   298
blanchet@41502
   299
(* FUDGE *)
blanchet@43517
   300
val smt_min_weight =
blanchet@43517
   301
  Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
blanchet@43517
   302
val smt_max_weight =
blanchet@43517
   303
  Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
blanchet@43517
   304
val smt_max_weight_index =
blanchet@43517
   305
  Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
blanchet@41502
   306
val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
blanchet@41502
   307
blanchet@43517
   308
fun smt_fact_weight ctxt j num_facts =
blanchet@43517
   309
  if Config.get ctxt smt_weights andalso
blanchet@43517
   310
     num_facts >= Config.get ctxt smt_weight_min_facts then
blanchet@43517
   311
    let
blanchet@43517
   312
      val min = Config.get ctxt smt_min_weight
blanchet@43517
   313
      val max = Config.get ctxt smt_max_weight
blanchet@43517
   314
      val max_index = Config.get ctxt smt_max_weight_index
blanchet@43517
   315
      val curve = !smt_weight_curve
blanchet@43517
   316
    in
blanchet@43517
   317
      SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
blanchet@43517
   318
            div curve max_index)
blanchet@43517
   319
    end
blanchet@41502
   320
  else
blanchet@41502
   321
    NONE
blanchet@41502
   322
blanchet@43517
   323
fun weight_smt_fact ctxt num_facts ((info, th), j) =
blanchet@43517
   324
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43517
   325
    (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
blanchet@43517
   326
  end
blanchet@38257
   327
blanchet@41339
   328
fun untranslated_fact (Untranslated_Fact p) = p
blanchet@41483
   329
  | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th)
blanchet@43450
   330
fun atp_translated_fact ctxt fact =
blanchet@43594
   331
  translate_atp_fact ctxt false (untranslated_fact fact)
blanchet@41502
   332
fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p
blanchet@43517
   333
  | smt_weighted_fact ctxt num_facts (fact, j) =
blanchet@43517
   334
    (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts
blanchet@41502
   335
blanchet@41561
   336
fun overlord_file_location_for_prover prover =
blanchet@41561
   337
  (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
blanchet@41561
   338
blanchet@41561
   339
blanchet@41502
   340
(* generic TPTP-based ATPs *)
blanchet@40242
   341
blanchet@43595
   342
(* Too general means, positive equality literal with a variable X as one
blanchet@43595
   343
   operand, when X does not occur properly in the other operand. This rules out
blanchet@43595
   344
   clearly inconsistent facts such as X = a | X = b, though it by no means
blanchet@43595
   345
   guarantees soundness. *)
blanchet@43595
   346
blanchet@43595
   347
(* Unwanted equalities are those between a (bound or schematic) variable that
blanchet@43595
   348
   does not properly occur in the second operand. *)
blanchet@43595
   349
val is_exhaustive_finite =
blanchet@43595
   350
  let
blanchet@43595
   351
    fun is_bad_equal (Var z) t =
blanchet@43595
   352
        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
blanchet@43595
   353
      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
blanchet@43595
   354
      | is_bad_equal _ _ = false
blanchet@43595
   355
    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
blanchet@43595
   356
    fun do_formula pos t =
blanchet@43595
   357
      case (pos, t) of
blanchet@43595
   358
        (_, @{const Trueprop} $ t1) => do_formula pos t1
blanchet@43595
   359
      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
blanchet@43595
   360
        do_formula pos t'
blanchet@43595
   361
      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
blanchet@43595
   362
        do_formula pos t'
blanchet@43595
   363
      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
blanchet@43595
   364
        do_formula pos t'
blanchet@43595
   365
      | (_, @{const "==>"} $ t1 $ t2) =>
blanchet@43595
   366
        do_formula (not pos) t1 andalso
blanchet@43595
   367
        (t2 = @{prop False} orelse do_formula pos t2)
blanchet@43595
   368
      | (_, @{const HOL.implies} $ t1 $ t2) =>
blanchet@43595
   369
        do_formula (not pos) t1 andalso
blanchet@43595
   370
        (t2 = @{const False} orelse do_formula pos t2)
blanchet@43595
   371
      | (_, @{const Not} $ t1) => do_formula (not pos) t1
blanchet@43595
   372
      | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
blanchet@43595
   373
      | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
blanchet@43595
   374
      | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
blanchet@43595
   375
      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
blanchet@43595
   376
      | _ => false
blanchet@43595
   377
  in do_formula true end
blanchet@43595
   378
blanchet@43595
   379
fun has_bound_or_var_of_type pred =
blanchet@43595
   380
  exists_subterm (fn Var (_, T as Type _) => pred T
blanchet@43595
   381
                   | Abs (_, T as Type _, _) => pred T
blanchet@43595
   382
                   | _ => false)
blanchet@43595
   383
blanchet@43595
   384
(* Facts are forbidden to contain variables of these types. The typical reason
blanchet@43595
   385
   is that they lead to unsoundness. Note that "unit" satisfies numerous
blanchet@43595
   386
   equations like "?x = ()". The resulting clauses will have no type constraint,
blanchet@43595
   387
   yielding false proofs. Even "bool" leads to many unsound proofs, though only
blanchet@43595
   388
   for higher-order problems. *)
blanchet@43595
   389
blanchet@43595
   390
(* Facts containing variables of type "unit" or "bool" or of the form
blanchet@43595
   391
   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
blanchet@43595
   392
   are omitted. *)
blanchet@43595
   393
fun is_dangerous_term ctxt =
blanchet@43595
   394
  transform_elim_term
blanchet@43595
   395
  #> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf
blanchet@43595
   396
      is_exhaustive_finite)
blanchet@43595
   397
blanchet@40636
   398
fun int_opt_add (SOME m) (SOME n) = SOME (m + n)
blanchet@40636
   399
  | int_opt_add _ _ = NONE
blanchet@40243
   400
blanchet@43323
   401
val atp_blacklist_max_iters = 10
blanchet@39732
   402
(* Important messages are important but not so important that users want to see
blanchet@39732
   403
   them each time. *)
blanchet@43480
   404
val atp_important_message_keep_quotient = 10
blanchet@39732
   405
blanchet@43460
   406
val fallback_best_type_systems =
blanchet@43708
   407
  [Preds (Polymorphic, Const_Arg_Types, Light),
blanchet@43708
   408
   Preds (Mangled_Monomorphic, Nonmonotonic_Types, Light)]
blanchet@43450
   409
blanchet@43587
   410
fun adjust_dumb_type_sys formats (Simple_Types level) =
blanchet@43587
   411
    if member (op =) formats Tff then (Tff, Simple_Types level)
blanchet@43708
   412
    else (Fof, Preds (Mangled_Monomorphic, level, Heavy))
blanchet@43718
   413
  | adjust_dumb_type_sys formats type_sys =
blanchet@43718
   414
    (* We could return (Tff, type_sys) in all cases but this would require the
blanchet@43718
   415
       TFF provers to accept problems in which constants are implicitly
blanchet@43718
   416
       declared. Today neither SNARK nor ToFoF-E satisfy this criterion. *)
blanchet@43718
   417
    if member (op =) formats Fof then (Fof, type_sys)
blanchet@43718
   418
    else (Tff, Simple_Types All_Types)
blanchet@43450
   419
fun determine_format_and_type_sys _ formats (Dumb_Type_Sys type_sys) =
blanchet@43450
   420
    adjust_dumb_type_sys formats type_sys
blanchet@43450
   421
  | determine_format_and_type_sys best_type_systems formats
blanchet@43450
   422
                                  (Smart_Type_Sys full_types) =
blanchet@43484
   423
    map type_sys_from_string best_type_systems @ fallback_best_type_systems
blanchet@43460
   424
    |> find_first (if full_types then is_type_sys_virtually_sound else K true)
blanchet@43460
   425
    |> the |> adjust_dumb_type_sys formats
blanchet@43419
   426
blanchet@43605
   427
fun repair_smt_monomorph_context debug max_mono_iters max_mono_instances =
blanchet@43589
   428
  Config.put SMT_Config.verbose debug
blanchet@43589
   429
  #> Config.put SMT_Config.monomorph_full false
blanchet@43589
   430
  #> Config.put SMT_Config.monomorph_limit max_mono_iters
blanchet@43605
   431
  #> Config.put SMT_Config.monomorph_instances max_mono_instances
blanchet@43589
   432
blanchet@43315
   433
fun run_atp auto name
blanchet@43449
   434
        ({exec, required_execs, arguments, proof_delims, known_failures,
blanchet@43588
   435
          conj_sym_kind, prem_kind, formats, best_slices, ...} : atp_config)
blanchet@43589
   436
        ({debug, verbose, overlord, type_sys, max_relevant, max_mono_iters,
blanchet@43605
   437
          max_new_mono_instances, explicit_apply, isar_proof,
blanchet@43605
   438
          isar_shrink_factor, slicing, timeout, ...} : params)
blanchet@41229
   439
        minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
blanchet@38257
   440
  let
blanchet@43053
   441
    val thy = Proof.theory_of state
blanchet@39564
   442
    val ctxt = Proof.context_of state
blanchet@39242
   443
    val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
blanchet@41407
   444
    val (dest_dir, problem_prefix) =
blanchet@41407
   445
      if overlord then overlord_file_location_for_prover name
blanchet@41407
   446
      else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
blanchet@40242
   447
    val problem_file_name =
blanchet@41407
   448
      Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
blanchet@41407
   449
                  "_" ^ string_of_int subgoal)
blanchet@40242
   450
    val problem_path_name =
blanchet@40240
   451
      if dest_dir = "" then
blanchet@40242
   452
        File.tmp_path problem_file_name
blanchet@40240
   453
      else if File.exists (Path.explode dest_dir) then
blanchet@40242
   454
        Path.append (Path.explode dest_dir) problem_file_name
blanchet@39247
   455
      else
blanchet@40240
   456
        error ("No such directory: " ^ quote dest_dir ^ ".")
blanchet@39247
   457
    val measure_run_time = verbose orelse Config.get ctxt measure_run_time
blanchet@38338
   458
    val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
blanchet@38257
   459
    fun split_time s =
blanchet@38257
   460
      let
blanchet@43319
   461
        val split = String.tokens (fn c => str c = "\n")
blanchet@43319
   462
        val (output, t) = s |> split |> split_last |> apfst cat_lines
blanchet@43319
   463
        fun as_num f = f >> (fst o read_int)
blanchet@43319
   464
        val num = as_num (Scan.many1 Symbol.is_ascii_digit)
blanchet@43319
   465
        val digit = Scan.one Symbol.is_ascii_digit
blanchet@43319
   466
        val num3 = as_num (digit ::: digit ::: (digit >> single))
blanchet@43319
   467
        val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
wenzelm@40875
   468
        val as_time = Scan.read Symbol.stopper time o raw_explode
blanchet@43319
   469
      in (output, as_time t) end
blanchet@41561
   470
    fun run_on prob_file =
blanchet@38338
   471
      case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
blanchet@38278
   472
        (home_var, _) :: _ =>
blanchet@38257
   473
        error ("The environment variable " ^ quote home_var ^ " is not set.")
blanchet@38278
   474
      | [] =>
blanchet@38278
   475
        if File.exists command then
blanchet@38278
   476
          let
blanchet@43314
   477
            (* If slicing is disabled, we expand the last slice to fill the
blanchet@43314
   478
               entire time available. *)
blanchet@43517
   479
            val actual_slices = get_slices slicing (best_slices ctxt)
blanchet@43314
   480
            val num_actual_slices = length actual_slices
blanchet@43316
   481
            fun monomorphize_facts facts =
blanchet@43316
   482
              let
blanchet@43316
   483
                val facts = facts |> map untranslated_fact
blanchet@43316
   484
                (* pseudo-theorem involving the same constants as the subgoal *)
blanchet@43316
   485
                val subgoal_th =
blanchet@43316
   486
                  Logic.list_implies (hyp_ts, concl_t)
blanchet@43316
   487
                  |> Skip_Proof.make_thm thy
blanchet@43316
   488
                val indexed_facts =
blanchet@43316
   489
                  (~1, subgoal_th) :: (0 upto length facts - 1 ~~ map snd facts)
blanchet@43605
   490
                val max_mono_instances = max_new_mono_instances + length facts
blanchet@43316
   491
              in
blanchet@43589
   492
                ctxt |> repair_smt_monomorph_context debug max_mono_iters
blanchet@43605
   493
                                                     max_mono_instances
blanchet@43589
   494
                     |> SMT_Monomorph.monomorph indexed_facts
blanchet@43589
   495
                     |> fst |> sort (int_ord o pairself fst)
blanchet@43589
   496
                     |> filter_out (curry (op =) ~1 o fst)
blanchet@43589
   497
                     |> map (Untranslated_Fact o apfst (fst o nth facts))
blanchet@43316
   498
              end
blanchet@43588
   499
            fun run_slice blacklist (slice, (time_frac, (complete,
blanchet@43588
   500
                                       (best_max_relevant, best_type_systems))))
blanchet@43314
   501
                          time_left =
blanchet@38278
   502
              let
blanchet@43314
   503
                val num_facts =
blanchet@43314
   504
                  length facts |> is_none max_relevant
blanchet@43588
   505
                                  ? Integer.min best_max_relevant
blanchet@43588
   506
                val (format, type_sys) =
blanchet@43588
   507
                  determine_format_and_type_sys best_type_systems formats type_sys
blanchet@43509
   508
                val fairly_sound = is_type_sys_fairly_sound type_sys
blanchet@43322
   509
                val facts =
blanchet@43532
   510
                  facts |> not fairly_sound
blanchet@43595
   511
                           ? filter_out (is_dangerous_term ctxt o prop_of o snd
blanchet@43509
   512
                                         o untranslated_fact)
blanchet@43509
   513
                        |> take num_facts
blanchet@43323
   514
                        |> not (null blacklist)
blanchet@43323
   515
                           ? filter_out (member (op =) blacklist o fst
blanchet@43323
   516
                                         o untranslated_fact)
blanchet@43460
   517
                        |> polymorphism_of_type_sys type_sys <> Polymorphic
blanchet@43460
   518
                           ? monomorphize_facts
blanchet@43415
   519
                        |> map (atp_translated_fact ctxt)
blanchet@43314
   520
                val real_ms = Real.fromInt o Time.toMilliseconds
blanchet@43314
   521
                val slice_timeout =
blanchet@43314
   522
                  ((real_ms time_left
blanchet@43314
   523
                    |> (if slice < num_actual_slices - 1 then
blanchet@43314
   524
                          curry Real.min (time_frac * real_ms timeout)
blanchet@43314
   525
                        else
blanchet@43314
   526
                          I))
blanchet@43314
   527
                   * 0.001) |> seconds
blanchet@43314
   528
                val _ =
blanchet@43485
   529
                  if debug then
blanchet@43569
   530
                    quote name ^ " slice #" ^ string_of_int (slice + 1) ^
blanchet@43569
   531
                    " with " ^ string_of_int num_facts ^ " fact" ^
blanchet@43569
   532
                    plural_s num_facts ^ " for " ^
blanchet@43569
   533
                    string_from_time slice_timeout ^ "..."
blanchet@43314
   534
                    |> Output.urgent_message
blanchet@43314
   535
                  else
blanchet@43314
   536
                    ()
blanchet@43412
   537
                val (atp_problem, pool, conjecture_offset, facts_offset,
blanchet@43620
   538
                     fact_names, sym_tab) =
blanchet@43580
   539
                  prepare_atp_problem ctxt conj_sym_kind prem_kind type_sys
blanchet@43580
   540
                                      explicit_apply hyp_ts concl_t facts
blanchet@41561
   541
                fun weights () = atp_problem_weights atp_problem
blanchet@41561
   542
                val core =
blanchet@41561
   543
                  File.shell_path command ^ " " ^
blanchet@43517
   544
                  arguments ctxt slice slice_timeout weights ^ " " ^
blanchet@41561
   545
                  File.shell_path prob_file
blanchet@41561
   546
                val command =
blanchet@41561
   547
                  (if measure_run_time then
blanchet@41561
   548
                     "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }"
blanchet@41561
   549
                   else
blanchet@41561
   550
                     "exec " ^ core) ^ " 2>&1"
blanchet@43314
   551
                val _ =
blanchet@43314
   552
                  atp_problem
blanchet@43580
   553
                  |> tptp_strings_for_atp_problem format
blanchet@43314
   554
                  |> cons ("% " ^ command ^ "\n")
blanchet@43314
   555
                  |> File.write_list prob_file
blanchet@43314
   556
                val conjecture_shape =
blanchet@43314
   557
                  conjecture_offset + 1
blanchet@43314
   558
                    upto conjecture_offset + length hyp_ts + 1
blanchet@43314
   559
                  |> map single
blanchet@38278
   560
                val ((output, msecs), res_code) =
blanchet@38278
   561
                  bash_output command
blanchet@38278
   562
                  |>> (if overlord then
blanchet@38278
   563
                         prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
blanchet@38278
   564
                       else
blanchet@38278
   565
                         I)
blanchet@40243
   566
                  |>> (if measure_run_time then split_time else rpair NONE)
blanchet@43320
   567
                val (atp_proof, outcome) =
blanchet@43718
   568
                  extract_tstplike_proof_and_outcome verbose complete res_code
blanchet@43718
   569
                      proof_delims known_failures output
blanchet@43320
   570
                  |>> atp_proof_from_tstplike_proof
blanchet@43458
   571
                val (conjecture_shape, facts_offset, fact_names) =
blanchet@43320
   572
                  if is_none outcome then
blanchet@43518
   573
                    repair_conjecture_shape_and_fact_names type_sys output
blanchet@43458
   574
                        conjecture_shape facts_offset fact_names
blanchet@43320
   575
                  else
blanchet@43458
   576
                    (* don't bother repairing *)
blanchet@43458
   577
                    (conjecture_shape, facts_offset, fact_names)
blanchet@43320
   578
                val outcome =
blanchet@43322
   579
                  case outcome of
blanchet@43458
   580
                    NONE =>
blanchet@43518
   581
                    if is_unsound_proof type_sys conjecture_shape facts_offset
blanchet@43518
   582
                                        fact_names atp_proof then
blanchet@43460
   583
                      SOME (UnsoundProof (is_type_sys_virtually_sound type_sys))
blanchet@43458
   584
                    else
blanchet@43458
   585
                      NONE
blanchet@43322
   586
                  | SOME Unprovable =>
blanchet@43323
   587
                    if null blacklist then outcome
blanchet@43322
   588
                    else SOME IncompleteUnprovable
blanchet@43322
   589
                  | _ => outcome
blanchet@43314
   590
              in
blanchet@43620
   591
                ((pool, conjecture_shape, facts_offset, fact_names, sym_tab),
blanchet@43588
   592
                 (output, msecs, type_sys, atp_proof, outcome))
blanchet@43314
   593
              end
blanchet@38883
   594
            val timer = Timer.startRealTimer ()
blanchet@43323
   595
            fun maybe_run_slice blacklist slice
blanchet@43588
   596
                                (result as (_, (_, msecs0, _, _, SOME _))) =
blanchet@43323
   597
                let
blanchet@43323
   598
                  val time_left = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@43323
   599
                in
blanchet@43323
   600
                  if Time.<= (time_left, Time.zeroTime) then
blanchet@43323
   601
                    result
blanchet@43323
   602
                  else
blanchet@43323
   603
                    (run_slice blacklist slice time_left
blanchet@43588
   604
                     |> (fn (stuff, (output, msecs, type_sys, atp_proof,
blanchet@43588
   605
                                     outcome)) =>
blanchet@43588
   606
                            (stuff, (output, int_opt_add msecs0 msecs, type_sys,
blanchet@43588
   607
                                     atp_proof, outcome))))
blanchet@43323
   608
                end
blanchet@43321
   609
              | maybe_run_slice _ _ result = result
blanchet@43323
   610
            fun maybe_blacklist_facts_and_retry iter blacklist
blanchet@43620
   611
                    (result as ((_, _, facts_offset, fact_names, _),
blanchet@43588
   612
                                (_, _, type_sys, atp_proof,
blanchet@43588
   613
                                 SOME (UnsoundProof false)))) =
blanchet@43518
   614
                (case used_facts_in_atp_proof type_sys facts_offset fact_names
blanchet@43412
   615
                                              atp_proof of
blanchet@43323
   616
                   [] => result
blanchet@43323
   617
                 | new_baddies =>
blanchet@43706
   618
                   if slicing andalso iter < atp_blacklist_max_iters - 1 then
blanchet@43648
   619
                     let val blacklist = new_baddies @ blacklist in
blanchet@43648
   620
                       result
blanchet@43648
   621
                       |> maybe_run_slice blacklist (List.last actual_slices)
blanchet@43648
   622
                       |> maybe_blacklist_facts_and_retry (iter + 1) blacklist
blanchet@43648
   623
                     end
blanchet@43648
   624
                   else
blanchet@43648
   625
                     result)
blanchet@43323
   626
              | maybe_blacklist_facts_and_retry _ _ result = result
blanchet@43314
   627
          in
blanchet@43620
   628
            ((Symtab.empty, [], 0, Vector.fromList [], Symtab.empty),
blanchet@43588
   629
             ("", SOME 0, hd fallback_best_type_systems, [],
blanchet@43588
   630
              SOME InternalError))
blanchet@43322
   631
            |> fold (maybe_run_slice []) actual_slices
blanchet@43322
   632
               (* The ATP found an unsound proof? Automatically try again
blanchet@43322
   633
                  without the offending facts! *)
blanchet@43323
   634
            |> maybe_blacklist_facts_and_retry 0 []
blanchet@43314
   635
          end
blanchet@38278
   636
        else
wenzelm@42815
   637
          error ("Bad executable: " ^ Path.print command ^ ".")
blanchet@38257
   638
blanchet@38257
   639
    (* If the problem file has not been exported, remove it; otherwise, export
blanchet@38257
   640
       the proof file too. *)
blanchet@41561
   641
    fun cleanup prob_file =
blanchet@41561
   642
      if dest_dir = "" then try File.rm prob_file else NONE
blanchet@43588
   643
    fun export prob_file (_, (output, _, _, _, _)) =
blanchet@40240
   644
      if dest_dir = "" then
blanchet@38257
   645
        ()
blanchet@38257
   646
      else
blanchet@41561
   647
        File.write (Path.explode (Path.implode prob_file ^ "_proof")) output
blanchet@43620
   648
    val ((pool, conjecture_shape, facts_offset, fact_names, sym_tab),
blanchet@43588
   649
         (output, msecs, type_sys, atp_proof, outcome)) =
blanchet@40242
   650
      with_path cleanup export run_on problem_path_name
blanchet@39732
   651
    val important_message =
blanchet@43480
   652
      if not auto andalso
blanchet@43480
   653
         random_range 0 (atp_important_message_keep_quotient - 1) = 0 then
blanchet@39732
   654
        extract_important_message output
blanchet@39732
   655
      else
blanchet@39732
   656
        ""
blanchet@42613
   657
    fun append_to_message message =
blanchet@42613
   658
      message ^
blanchet@42613
   659
      (if verbose then
blanchet@42613
   660
         "\nATP real CPU time: " ^
blanchet@42613
   661
         string_from_time (Time.fromMilliseconds (the msecs)) ^ "."
blanchet@42613
   662
       else
blanchet@42613
   663
         "") ^
blanchet@42613
   664
      (if important_message <> "" then
blanchet@42613
   665
         "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message
blanchet@42613
   666
       else
blanchet@42613
   667
         "")
blanchet@43620
   668
    val isar_params =
blanchet@43620
   669
      (ctxt, debug, isar_shrink_factor, pool, conjecture_shape, sym_tab)
blanchet@42613
   670
    val metis_params =
blanchet@43518
   671
      (proof_banner auto, minimize_command, type_sys, atp_proof, facts_offset,
blanchet@43518
   672
       fact_names, goal, subgoal)
blanchet@42613
   673
    val (outcome, (message, used_facts)) =
blanchet@38257
   674
      case outcome of
blanchet@38257
   675
        NONE =>
blanchet@42613
   676
        (NONE, proof_text isar_proof isar_params metis_params
blanchet@42613
   677
               |>> append_to_message)
blanchet@42613
   678
      | SOME ProofMissing =>
blanchet@42613
   679
        (NONE, metis_proof_text metis_params |>> append_to_message)
blanchet@42615
   680
      | SOME failure => (outcome, (string_for_failure failure, []))
blanchet@38257
   681
  in
blanchet@40445
   682
    {outcome = outcome, message = message, used_facts = used_facts,
blanchet@40242
   683
     run_time_in_msecs = msecs}
blanchet@38257
   684
  end
blanchet@38257
   685
blanchet@40917
   686
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
blanchet@40917
   687
   these are sorted out properly in the SMT module, we have to interpret these
blanchet@40917
   688
   ourselves. *)
blanchet@40932
   689
val remote_smt_failures =
blanchet@40932
   690
  [(22, CantConnect),
blanchet@40932
   691
   (2, NoLibwwwPerl)]
blanchet@41470
   692
val z3_wrapper_failures =
blanchet@41470
   693
  [(10, NoRealZ3),
blanchet@41470
   694
   (11, InternalError),
blanchet@41470
   695
   (12, InternalError),
blanchet@41470
   696
   (13, InternalError)]
blanchet@40932
   697
val z3_failures =
blanchet@41477
   698
  [(101, OutOfResources),
blanchet@41477
   699
   (103, MalformedInput),
blanchet@41470
   700
   (110, MalformedInput)]
blanchet@40932
   701
val unix_failures =
blanchet@40932
   702
  [(139, Crashed)]
blanchet@41470
   703
val smt_failures =
blanchet@42670
   704
  remote_smt_failures @ z3_wrapper_failures @ z3_failures @ unix_failures
blanchet@40797
   705
blanchet@42964
   706
fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
blanchet@42964
   707
    if is_real_cex then Unprovable else IncompleteUnprovable
blanchet@41470
   708
  | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
blanchet@41470
   709
  | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
blanchet@41470
   710
    (case AList.lookup (op =) smt_failures code of
blanchet@40932
   711
       SOME failure => failure
blanchet@41505
   712
     | NONE => UnknownError ("Abnormal termination with exit code " ^
blanchet@41505
   713
                             string_of_int code ^ "."))
blanchet@41470
   714
  | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
blanchet@41470
   715
  | failure_from_smt_failure (SMT_Failure.Other_Failure msg) =
blanchet@42916
   716
    UnknownError msg
wenzelm@28586
   717
blanchet@40946
   718
(* FUDGE *)
blanchet@43517
   719
val smt_max_slices =
blanchet@43517
   720
  Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
blanchet@43517
   721
val smt_slice_fact_frac =
blanchet@43517
   722
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
blanchet@43517
   723
val smt_slice_time_frac =
blanchet@43517
   724
  Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
blanchet@43517
   725
val smt_slice_min_secs =
blanchet@43517
   726
  Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
blanchet@40636
   727
blanchet@43517
   728
fun smt_filter_loop ctxt name
blanchet@43589
   729
                    ({debug, verbose, overlord, max_mono_iters,
blanchet@43605
   730
                      max_new_mono_instances, timeout, slicing, ...} : params)
blanchet@42612
   731
                    state i smt_filter =
blanchet@40636
   732
  let
blanchet@43517
   733
    val max_slices = if slicing then Config.get ctxt smt_max_slices else 1
blanchet@41483
   734
    val repair_context =
blanchet@43589
   735
          select_smt_solver name
blanchet@43589
   736
          #> (if overlord then
blanchet@43589
   737
                Config.put SMT_Config.debug_files
blanchet@43589
   738
                           (overlord_file_location_for_prover name
blanchet@43589
   739
                            |> (fn (path, name) => path ^ "/" ^ name))
blanchet@43589
   740
              else
blanchet@43589
   741
                I)
blanchet@43589
   742
          #> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers)
blanchet@41483
   743
    val state = state |> Proof.map_context repair_context
blanchet@43314
   744
    fun do_slice timeout slice outcome0 time_so_far facts =
blanchet@40795
   745
      let
blanchet@40795
   746
        val timer = Timer.startRealTimer ()
blanchet@43589
   747
        val state =
blanchet@43589
   748
          state |> Proof.map_context
blanchet@43589
   749
                       (repair_smt_monomorph_context debug max_mono_iters
blanchet@43605
   750
                            (max_new_mono_instances + length facts))
blanchet@40795
   751
        val ms = timeout |> Time.toMilliseconds
blanchet@43314
   752
        val slice_timeout =
blanchet@43314
   753
          if slice < max_slices then
blanchet@41415
   754
            Int.min (ms,
blanchet@43517
   755
                Int.max (1000 * Config.get ctxt smt_slice_min_secs,
blanchet@43517
   756
                    Real.ceil (Config.get ctxt smt_slice_time_frac
blanchet@43517
   757
                               * Real.fromInt ms)))
blanchet@40795
   758
            |> Time.fromMilliseconds
blanchet@40795
   759
          else
blanchet@40795
   760
            timeout
blanchet@40795
   761
        val num_facts = length facts
blanchet@40795
   762
        val _ =
blanchet@43485
   763
          if debug then
blanchet@43485
   764
            quote name ^ " slice " ^ string_of_int slice ^ " with " ^
blanchet@43485
   765
            string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^
blanchet@43485
   766
            string_from_time slice_timeout ^ "..."
blanchet@40795
   767
            |> Output.urgent_message
blanchet@40795
   768
          else
blanchet@40795
   769
            ()
blanchet@41414
   770
        val birth = Timer.checkRealTimer timer
blanchet@41419
   771
        val _ =
blanchet@41459
   772
          if debug then Output.urgent_message "Invoking SMT solver..." else ()
blanchet@41457
   773
        val (outcome, used_facts) =
blanchet@43314
   774
          (case (slice, smt_filter) of
boehmes@41680
   775
             (1, SOME head) => head |> apsnd (apsnd repair_context)
boehmes@41680
   776
           | _ => SMT_Solver.smt_filter_preprocess state facts i)
blanchet@43314
   777
          |> SMT_Solver.smt_filter_apply slice_timeout
blanchet@41480
   778
          |> (fn {outcome, used_facts} => (outcome, used_facts))
blanchet@41457
   779
          handle exn => if Exn.is_interrupt exn then
blanchet@41457
   780
                          reraise exn
blanchet@41457
   781
                        else
blanchet@42916
   782
                          (ML_Compiler.exn_message exn
blanchet@41457
   783
                           |> SMT_Failure.Other_Failure |> SOME, [])
blanchet@41414
   784
        val death = Timer.checkRealTimer timer
blanchet@40795
   785
        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
blanchet@41414
   786
        val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
blanchet@40795
   787
        val too_many_facts_perhaps =
blanchet@40795
   788
          case outcome of
blanchet@40795
   789
            NONE => false
blanchet@40795
   790
          | SOME (SMT_Failure.Counterexample _) => false
blanchet@43314
   791
          | SOME SMT_Failure.Time_Out => slice_timeout <> timeout
blanchet@43485
   792
          | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
blanchet@40795
   793
          | SOME SMT_Failure.Out_Of_Memory => true
blanchet@41459
   794
          | SOME (SMT_Failure.Other_Failure _) => true
blanchet@40795
   795
        val timeout = Time.- (timeout, Timer.checkRealTimer timer)
blanchet@40795
   796
      in
blanchet@43314
   797
        if too_many_facts_perhaps andalso slice < max_slices andalso
blanchet@40795
   798
           num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
blanchet@41415
   799
          let
blanchet@43485
   800
            val new_num_facts =
blanchet@43517
   801
              Real.ceil (Config.get ctxt smt_slice_fact_frac
blanchet@43517
   802
                         * Real.fromInt num_facts)
blanchet@43485
   803
            val _ =
blanchet@43485
   804
              if verbose andalso is_some outcome then
blanchet@43485
   805
                quote name ^ " invoked with " ^ string_of_int num_facts ^
blanchet@43485
   806
                " fact" ^ plural_s num_facts ^ ": " ^
blanchet@43485
   807
                string_for_failure (failure_from_smt_failure (the outcome)) ^
blanchet@43485
   808
                " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^
blanchet@43509
   809
                plural_s new_num_facts ^ "..."
blanchet@43485
   810
                |> Output.urgent_message
blanchet@43485
   811
              else
blanchet@43485
   812
                ()
blanchet@43314
   813
          in
blanchet@43485
   814
            facts |> take new_num_facts
blanchet@43485
   815
                  |> do_slice timeout (slice + 1) outcome0 time_so_far
blanchet@43314
   816
          end
blanchet@40795
   817
        else
blanchet@40795
   818
          {outcome = if is_none outcome then NONE else the outcome0,
blanchet@41414
   819
           used_facts = used_facts,
blanchet@41414
   820
           run_time_in_msecs = SOME (Time.toMilliseconds time_so_far)}
blanchet@40636
   821
      end
blanchet@43314
   822
  in do_slice timeout 1 NONE Time.zeroTime end
blanchet@40636
   823
blanchet@40914
   824
(* taken from "Mirabelle" and generalized *)
blanchet@40914
   825
fun can_apply timeout tac state i =
blanchet@40914
   826
  let
blanchet@40914
   827
    val {context = ctxt, facts, goal} = Proof.goal state
blanchet@40914
   828
    val full_tac = Method.insert_tac facts i THEN tac ctxt i
blanchet@40914
   829
  in
blanchet@40914
   830
    case try (TimeLimit.timeLimit timeout (Seq.pull o full_tac)) goal of
blanchet@40914
   831
      SOME (SOME _) => true
blanchet@40914
   832
    | _ => false
blanchet@40914
   833
  end
blanchet@40914
   834
blanchet@41400
   835
val smt_metis_timeout = seconds 1.0
blanchet@40914
   836
blanchet@40941
   837
fun can_apply_metis debug state i ths =
blanchet@40941
   838
  can_apply smt_metis_timeout
blanchet@40941
   839
            (Config.put Metis_Tactics.verbose debug
blanchet@40941
   840
             #> (fn ctxt => Metis_Tactics.metis_tac ctxt ths)) state i
blanchet@40914
   841
blanchet@43315
   842
fun run_smt_solver auto name (params as {debug, verbose, ...}) minimize_command
blanchet@43314
   843
                   ({state, subgoal, subgoal_count, facts, smt_filter, ...}
blanchet@43314
   844
                    : prover_problem) =
blanchet@36379
   845
  let
blanchet@41483
   846
    val ctxt = Proof.context_of state
blanchet@41502
   847
    val num_facts = length facts
blanchet@41502
   848
    val facts = facts ~~ (0 upto num_facts - 1)
blanchet@43517
   849
                |> map (smt_weighted_fact ctxt num_facts)
blanchet@40422
   850
    val {outcome, used_facts, run_time_in_msecs} =
blanchet@43517
   851
      smt_filter_loop ctxt name params state subgoal smt_filter facts
blanchet@40966
   852
    val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
blanchet@41470
   853
    val outcome = outcome |> Option.map failure_from_smt_failure
blanchet@40244
   854
    val message =
blanchet@40425
   855
      case outcome of
blanchet@40425
   856
        NONE =>
blanchet@40914
   857
        let
blanchet@41399
   858
          val (method, settings) =
blanchet@40941
   859
            if can_apply_metis debug state subgoal (map snd used_facts) then
blanchet@41399
   860
              ("metis", "")
blanchet@40941
   861
            else
boehmes@41680
   862
              ("smt", if name = SMT_Solver.solver_name_of ctxt then ""
boehmes@41680
   863
                      else "smt_solver = " ^ maybe_quote name)
blanchet@40914
   864
        in
blanchet@40914
   865
          try_command_line (proof_banner auto)
blanchet@41399
   866
              (apply_on_subgoal settings subgoal subgoal_count ^
blanchet@41399
   867
               command_call method (map fst other_lemmas)) ^
blanchet@40966
   868
          minimize_line minimize_command
blanchet@41414
   869
                        (map fst (other_lemmas @ chained_lemmas)) ^
blanchet@41414
   870
          (if verbose then
blanchet@41414
   871
             "\nSMT solver real CPU time: " ^
blanchet@41414
   872
             string_from_time (Time.fromMilliseconds (the run_time_in_msecs)) ^
blanchet@41414
   873
             "."
blanchet@41414
   874
           else
blanchet@41414
   875
             "")
blanchet@40914
   876
        end
blanchet@42615
   877
      | SOME failure => string_for_failure failure
blanchet@40244
   878
  in
blanchet@40914
   879
    {outcome = outcome, used_facts = map fst used_facts,
blanchet@40445
   880
     run_time_in_msecs = run_time_in_msecs, message = message}
blanchet@40244
   881
  end
blanchet@40244
   882
blanchet@43315
   883
fun get_prover ctxt auto name =
wenzelm@43232
   884
  let val thy = Proof_Context.theory_of ctxt in
blanchet@41189
   885
    if is_smt_prover ctxt name then
blanchet@43315
   886
      run_smt_solver auto name
blanchet@42591
   887
    else if member (op =) (supported_atps thy) name then
blanchet@43315
   888
      run_atp auto name (get_atp thy name)
blanchet@41189
   889
    else
blanchet@41189
   890
      error ("No such prover: " ^ name ^ ".")
blanchet@41189
   891
  end
blanchet@40244
   892
wenzelm@28582
   893
end;