src/HOL/Tools/ATP_Manager/atp_wrapper.ML
author blanchet
Wed, 14 Apr 2010 18:23:51 +0200
changeset 36140 f5e15e9aae10
parent 36064 48aec67c284f
child 36141 6490319b1703
permissions -rw-r--r--
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
wenzelm@32327
     1
(*  Title:      HOL/Tools/ATP_Manager/atp_wrapper.ML
wenzelm@28592
     2
    Author:     Fabian Immler, TU Muenchen
wenzelm@28592
     3
wenzelm@28592
     4
Wrapper functions for external ATPs.
wenzelm@28592
     5
*)
wenzelm@28592
     6
wenzelm@28592
     7
signature ATP_WRAPPER =
wenzelm@28592
     8
sig
blanchet@35867
     9
  type prover = ATP_Manager.prover
boehmes@32864
    10
blanchet@35867
    11
  (* hooks for problem files *)
blanchet@35867
    12
  val destdir : string Config.T
blanchet@35867
    13
  val problem_prefix : string Config.T
blanchet@35867
    14
  val measure_runtime : bool Config.T
boehmes@32864
    15
blanchet@35867
    16
  val refresh_systems_on_tptp : unit -> unit
blanchet@35867
    17
  val setup : theory -> theory
wenzelm@28592
    18
end;
wenzelm@28592
    19
blanchet@35865
    20
structure ATP_Wrapper : ATP_WRAPPER =
wenzelm@28592
    21
struct
wenzelm@28596
    22
blanchet@35967
    23
open Sledgehammer_Fact_Preprocessor
blanchet@35865
    24
open Sledgehammer_HOL_Clause
blanchet@35865
    25
open Sledgehammer_Fact_Filter
blanchet@35865
    26
open Sledgehammer_Proof_Reconstruct
blanchet@35867
    27
open ATP_Manager
blanchet@35826
    28
wenzelm@28596
    29
(** generic ATP wrapper **)
wenzelm@28596
    30
wenzelm@32944
    31
(* external problem files *)
wenzelm@28596
    32
wenzelm@36001
    33
val (destdir, destdir_setup) = Attrib.config_string "atp_destdir" (K "");
boehmes@32864
    34
  (*Empty string means create files in Isabelle's temporary files directory.*)
boehmes@32864
    35
boehmes@32864
    36
val (problem_prefix, problem_prefix_setup) =
wenzelm@36001
    37
  Attrib.config_string "atp_problem_prefix" (K "prob");
boehmes@32864
    38
boehmes@33239
    39
val (measure_runtime, measure_runtime_setup) =
wenzelm@36001
    40
  Attrib.config_bool "atp_measure_runtime" (K false);
boehmes@33239
    41
boehmes@32864
    42
blanchet@35867
    43
(* prover configuration *)
boehmes@32864
    44
wenzelm@32941
    45
type prover_config =
wenzelm@32941
    46
 {command: Path.T,
blanchet@35967
    47
  arguments: Time.time -> string,
blanchet@35865
    48
  failure_strs: string list,
boehmes@32864
    49
  max_new_clauses: int,
blanchet@36059
    50
  prefers_theory_const: bool,
blanchet@35967
    51
  supports_isar_proofs: bool};
boehmes@32864
    52
wenzelm@28596
    53
wenzelm@28596
    54
(* basic template *)
wenzelm@28596
    55
boehmes@32458
    56
fun with_path cleanup after f path =
boehmes@32458
    57
  Exn.capture f path
boehmes@32458
    58
  |> tap (fn _ => cleanup path)
boehmes@32458
    59
  |> Exn.release
wenzelm@32941
    60
  |> tap (after path);
boehmes@32458
    61
blanchet@35865
    62
fun find_failure strs proof =
blanchet@35865
    63
  case filter (fn s => String.isSubstring s proof) strs of
blanchet@35865
    64
    [] => if is_proof_well_formed proof then NONE
blanchet@35865
    65
          else SOME "Ill-formed ATP output"
blanchet@35865
    66
  | (failure :: _) => SOME failure
blanchet@35865
    67
blanchet@35967
    68
fun generic_prover get_facts prepare write cmd args failure_strs produce_answer
blanchet@35967
    69
        name ({full_types, ...} : params)
blanchet@35967
    70
        ({subgoal, goal, relevance_override, axiom_clauses, filtered_clauses}
blanchet@35967
    71
         : problem) =
wenzelm@28596
    72
  let
immler@31750
    73
    (* get clauses and prepare them for writing *)
wenzelm@32942
    74
    val (ctxt, (chain_ths, th)) = goal;
wenzelm@32942
    75
    val thy = ProofContext.theory_of ctxt;
blanchet@35865
    76
    val chain_ths = map (Thm.put_name_hint chained_hint) chain_ths;
blanchet@35967
    77
    val goal_cls = #1 (neg_conjecture_clauses ctxt th subgoal);
immler@31752
    78
    val the_filtered_clauses =
wenzelm@32942
    79
      (case filtered_clauses of
blanchet@35967
    80
        NONE => get_facts relevance_override goal goal_cls
wenzelm@32942
    81
      | SOME fcls => fcls);
immler@31409
    82
    val the_axiom_clauses =
wenzelm@32942
    83
      (case axiom_clauses of
wenzelm@32942
    84
        NONE => the_filtered_clauses
wenzelm@32942
    85
      | SOME axcls => axcls);
blanchet@35967
    86
    val (internal_thm_names, clauses) =
wenzelm@32942
    87
      prepare goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy;
immler@31750
    88
boehmes@32864
    89
    (* path to unique problem file *)
wenzelm@32942
    90
    val destdir' = Config.get ctxt destdir;
wenzelm@32942
    91
    val problem_prefix' = Config.get ctxt problem_prefix;
boehmes@32864
    92
    fun prob_pathname nr =
wenzelm@32942
    93
      let val probfile =
wenzelm@32942
    94
        Path.basic (problem_prefix' ^ serial_string () ^ "_" ^ string_of_int nr)
wenzelm@32942
    95
      in
wenzelm@32942
    96
        if destdir' = "" then File.tmp_path probfile
wenzelm@35570
    97
        else if File.exists (Path.explode destdir')
wenzelm@35570
    98
        then Path.append  (Path.explode destdir') probfile
boehmes@32864
    99
        else error ("No such directory: " ^ destdir')
wenzelm@32942
   100
      end;
boehmes@32864
   101
immler@31750
   102
    (* write out problem file and call prover *)
boehmes@33239
   103
    fun cmd_line probfile =
boehmes@33239
   104
      if Config.get ctxt measure_runtime
boehmes@33239
   105
      then (* Warning: suppresses error messages of ATPs *)
boehmes@33239
   106
        "TIMEFORMAT='%3U'; { time " ^ space_implode " " [File.shell_path cmd,
boehmes@33239
   107
        args, File.shell_path probfile] ^ " 2> /dev/null" ^ " ; } 2>&1"
boehmes@33239
   108
      else
boehmes@33239
   109
        space_implode " " ["exec", File.shell_path cmd, args,
boehmes@33239
   110
        File.shell_path probfile];
boehmes@32510
   111
    fun split_time s =
boehmes@32510
   112
      let
wenzelm@32942
   113
        val split = String.tokens (fn c => str c = "\n");
wenzelm@32942
   114
        val (proof, t) = s |> split |> split_last |> apfst cat_lines;
wenzelm@32942
   115
        fun as_num f = f >> (fst o read_int);
wenzelm@32942
   116
        val num = as_num (Scan.many1 Symbol.is_ascii_digit);
wenzelm@32942
   117
        val digit = Scan.one Symbol.is_ascii_digit;
wenzelm@32942
   118
        val num3 = as_num (digit ::: digit ::: (digit >> single));
wenzelm@32942
   119
        val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b);
wenzelm@32942
   120
        val as_time = the_default 0 o Scan.read Symbol.stopper time o explode;
wenzelm@32942
   121
      in (proof, as_time t) end;
boehmes@33239
   122
    fun split_time' s =
boehmes@33239
   123
      if Config.get ctxt measure_runtime then split_time s else (s, 0)
boehmes@32458
   124
    fun run_on probfile =
wenzelm@32942
   125
      if File.exists cmd then
blanchet@35967
   126
        write full_types probfile clauses
wenzelm@35010
   127
        |> pair (apfst split_time' (bash_output (cmd_line probfile)))
wenzelm@32942
   128
      else error ("Bad executable: " ^ Path.implode cmd);
wenzelm@28592
   129
immler@31751
   130
    (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
wenzelm@32942
   131
    fun cleanup probfile = if destdir' = "" then try File.rm probfile else NONE;
wenzelm@32942
   132
    fun export probfile (((proof, _), _), _) =
wenzelm@32942
   133
      if destdir' = "" then ()
wenzelm@32942
   134
      else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof;
wenzelm@32261
   135
blanchet@35967
   136
    val (((proof, atp_run_time_in_msecs), rc), conj_pos) =
wenzelm@35570
   137
      with_path cleanup export run_on (prob_pathname subgoal);
boehmes@32458
   138
immler@29590
   139
    (* check for success and print out some information on failure *)
blanchet@35865
   140
    val failure = find_failure failure_strs proof;
wenzelm@32942
   141
    val success = rc = 0 andalso is_none failure;
blanchet@35967
   142
    val (message, relevant_thm_names) =
boehmes@32451
   143
      if is_some failure then ("External prover failed.", [])
boehmes@32451
   144
      else if rc <> 0 then ("External prover failed: " ^ proof, [])
blanchet@35967
   145
      else
blanchet@35967
   146
        (produce_answer name (proof, internal_thm_names, conj_pos, ctxt, th,
blanchet@35967
   147
                              subgoal));
boehmes@32864
   148
  in
wenzelm@32941
   149
    {success = success, message = message,
blanchet@35967
   150
     relevant_thm_names = relevant_thm_names,
blanchet@35967
   151
     atp_run_time_in_msecs = atp_run_time_in_msecs, proof = proof,
blanchet@35967
   152
     internal_thm_names = internal_thm_names,
blanchet@35967
   153
     filtered_clauses = the_filtered_clauses}
wenzelm@32942
   154
  end;
wenzelm@28596
   155
wenzelm@28596
   156
wenzelm@28596
   157
(* generic TPTP-based provers *)
wenzelm@28596
   158
blanchet@35865
   159
fun generic_tptp_prover
blanchet@35865
   160
        (name, {command, arguments, failure_strs, max_new_clauses,
blanchet@36059
   161
                prefers_theory_const, supports_isar_proofs})
blanchet@36058
   162
        (params as {respect_no_atp, relevance_threshold, convergence,
blanchet@36064
   163
                    theory_const, higher_order, follow_defs, isar_proof,
blanchet@36064
   164
                    modulus, sorts, ...})
blanchet@36059
   165
        timeout =
blanchet@35967
   166
  generic_prover
blanchet@36058
   167
      (get_relevant_facts respect_no_atp relevance_threshold convergence
blanchet@36058
   168
                          higher_order follow_defs max_new_clauses
blanchet@36059
   169
                          (the_default prefers_theory_const theory_const))
blanchet@35967
   170
      (prepare_clauses higher_order false) write_tptp_file command
blanchet@35967
   171
      (arguments timeout) failure_strs
blanchet@36064
   172
      (if supports_isar_proofs andalso isar_proof then
blanchet@36064
   173
         structured_isar_proof modulus sorts
blanchet@36064
   174
       else
blanchet@36064
   175
         metis_lemma_list false) name params;
wenzelm@28592
   176
blanchet@35967
   177
fun tptp_prover name p = (name, generic_tptp_prover (name, p));
wenzelm@28592
   178
wenzelm@32941
   179
boehmes@32864
   180
(** common provers **)
wenzelm@28596
   181
blanchet@36140
   182
fun generous_to_secs time = (Time.toMilliseconds time + 999) div 1000
blanchet@36140
   183
wenzelm@28596
   184
(* Vampire *)
wenzelm@28596
   185
blanchet@35967
   186
(* NB: Vampire does not work without explicit time limit. *)
wenzelm@28596
   187
blanchet@35967
   188
val vampire_config : prover_config =
blanchet@35967
   189
  {command = Path.explode "$VAMPIRE_HOME/vampire",
blanchet@35967
   190
   arguments = (fn timeout => "--output_syntax tptp --mode casc -t " ^
blanchet@36140
   191
                              string_of_int (generous_to_secs timeout)),
blanchet@35967
   192
   failure_strs =
blanchet@35967
   193
     ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"],
blanchet@35967
   194
   max_new_clauses = 60,
blanchet@36059
   195
   prefers_theory_const = false,
blanchet@35967
   196
   supports_isar_proofs = true}
blanchet@35967
   197
val vampire = tptp_prover "vampire" vampire_config
wenzelm@28596
   198
wenzelm@28596
   199
wenzelm@28596
   200
(* E prover *)
wenzelm@28596
   201
blanchet@35967
   202
val e_config : prover_config =
blanchet@35967
   203
  {command = Path.explode "$E_HOME/eproof",
blanchet@35967
   204
   arguments = (fn timeout => "--tstp-in --tstp-out -l5 -xAutoDev \
blanchet@35967
   205
                              \-tAutoDev --silent --cpu-limit=" ^
blanchet@36140
   206
                              string_of_int (generous_to_secs timeout)),
blanchet@35967
   207
   failure_strs =
blanchet@35967
   208
       ["SZS status: Satisfiable", "SZS status Satisfiable",
blanchet@35967
   209
        "SZS status: ResourceOut", "SZS status ResourceOut",
blanchet@35967
   210
        "# Cannot determine problem status"],
blanchet@35967
   211
   max_new_clauses = 100,
blanchet@36059
   212
   prefers_theory_const = false,
blanchet@35967
   213
   supports_isar_proofs = true}
blanchet@35967
   214
val e = tptp_prover "e" e_config
wenzelm@28596
   215
wenzelm@28596
   216
wenzelm@28596
   217
(* SPASS *)
wenzelm@28596
   218
blanchet@35865
   219
fun generic_dfg_prover
blanchet@35865
   220
        (name, ({command, arguments, failure_strs, max_new_clauses,
blanchet@36059
   221
                 prefers_theory_const, ...} : prover_config))
blanchet@36058
   222
        (params as {respect_no_atp, relevance_threshold, convergence,
blanchet@36064
   223
                    theory_const, higher_order, follow_defs, ...})
blanchet@36059
   224
        timeout =
blanchet@35967
   225
  generic_prover
blanchet@36058
   226
      (get_relevant_facts respect_no_atp relevance_threshold convergence
blanchet@36058
   227
                          higher_order follow_defs max_new_clauses
blanchet@36059
   228
                          (the_default prefers_theory_const theory_const))
blanchet@35967
   229
      (prepare_clauses higher_order true) write_dfg_file command
blanchet@35967
   230
      (arguments timeout) failure_strs (metis_lemma_list true) name params;
boehmes@32869
   231
blanchet@35865
   232
fun dfg_prover (name, p) = (name, generic_dfg_prover (name, p));
boehmes@32869
   233
blanchet@36059
   234
val spass_config : prover_config =
blanchet@35967
   235
 {command = Path.explode "$SPASS_HOME/SPASS",
blanchet@35967
   236
  arguments = (fn timeout => "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0" ^
blanchet@35967
   237
    " -FullRed=0 -DocProof -TimeLimit=" ^
blanchet@36140
   238
    string_of_int (generous_to_secs timeout)),
blanchet@35967
   239
  failure_strs =
blanchet@35967
   240
    ["SPASS beiseite: Completion found.", "SPASS beiseite: Ran out of time.",
blanchet@35967
   241
     "SPASS beiseite: Maximal number of loops exceeded."],
blanchet@35967
   242
  max_new_clauses = 40,
blanchet@36059
   243
  prefers_theory_const = true,
blanchet@36059
   244
  supports_isar_proofs = false}
blanchet@35967
   245
val spass = dfg_prover ("spass", spass_config)
blanchet@35967
   246
wenzelm@28596
   247
wenzelm@28596
   248
(* remote prover invocation via SystemOnTPTP *)
wenzelm@28596
   249
wenzelm@32942
   250
val systems = Synchronized.var "atp_wrapper_systems" ([]: string list);
immler@31828
   251
immler@31828
   252
fun get_systems () =
immler@31828
   253
  let
blanchet@35967
   254
    val (answer, rc) = bash_output "\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w"
immler@31828
   255
  in
blanchet@35967
   256
    if rc <> 0 then
blanchet@35967
   257
      error ("Failed to get available systems at SystemOnTPTP:\n" ^ answer)
blanchet@35967
   258
    else
blanchet@35967
   259
      split_lines answer
immler@31828
   260
  end;
immler@31828
   261
blanchet@35867
   262
fun refresh_systems_on_tptp () =
blanchet@35867
   263
  Synchronized.change systems (fn _ => get_systems ());
immler@31828
   264
immler@31828
   265
fun get_system prefix = Synchronized.change_result systems (fn systems =>
boehmes@32864
   266
  (if null systems then get_systems () else systems)
wenzelm@32942
   267
  |> `(find_first (String.isPrefix prefix)));
immler@31828
   268
wenzelm@32948
   269
fun the_system prefix =
boehmes@32864
   270
  (case get_system prefix of
blanchet@35826
   271
    NONE => error ("System " ^ quote prefix ^ " not available at SystemOnTPTP")
wenzelm@32942
   272
  | SOME sys => sys);
wenzelm@28596
   273
blanchet@35865
   274
val remote_failure_strs = ["Remote-script could not extract proof"];
blanchet@35865
   275
blanchet@36059
   276
fun remote_prover_config prover_prefix args
blanchet@36059
   277
        ({failure_strs, max_new_clauses, prefers_theory_const, ...}
blanchet@36059
   278
         : prover_config) : prover_config =
blanchet@35967
   279
  {command = Path.explode "$ISABELLE_ATP_MANAGER/SystemOnTPTP",
blanchet@35967
   280
   arguments = (fn timeout =>
blanchet@36140
   281
     args ^ " -t " ^ string_of_int (generous_to_secs timeout) ^ " -s " ^
blanchet@35967
   282
     the_system prover_prefix),
blanchet@36059
   283
   failure_strs = remote_failure_strs @ failure_strs,
blanchet@35967
   284
   max_new_clauses = max_new_clauses,
blanchet@36059
   285
   prefers_theory_const = prefers_theory_const,
blanchet@35967
   286
   supports_isar_proofs = false}
boehmes@32864
   287
blanchet@35967
   288
val remote_vampire =
blanchet@35967
   289
  tptp_prover "remote_vampire"
blanchet@36059
   290
              (remote_prover_config "Vampire---9" "" vampire_config)
boehmes@32864
   291
blanchet@35967
   292
val remote_e =
blanchet@36059
   293
  tptp_prover "remote_e" (remote_prover_config "EP---" "" e_config)
boehmes@32864
   294
blanchet@35967
   295
val remote_spass =
blanchet@36059
   296
  tptp_prover "remote_spass" (remote_prover_config "SPASS---" "-x" spass_config)
wenzelm@28592
   297
blanchet@36059
   298
val provers = [spass, vampire, e, remote_vampire, remote_spass, remote_e]
blanchet@35867
   299
val prover_setup = fold add_prover provers
blanchet@35867
   300
blanchet@35867
   301
val setup =
blanchet@35867
   302
  destdir_setup
blanchet@35867
   303
  #> problem_prefix_setup
blanchet@35867
   304
  #> measure_runtime_setup
blanchet@35867
   305
  #> prover_setup;
blanchet@35867
   306
wenzelm@28592
   307
end;