src/HOL/Tools/ATP/atp_systems.ML
author blanchet
Tue, 09 Aug 2011 17:33:17 +0200
changeset 44967 6e943b3d2767
parent 44963 bf489e54d7f8
child 44970 5e14f591515e
permissions -rw-r--r--
LEO-II also supports FOF
blanchet@38293
     1
(*  Title:      HOL/Tools/ATP/atp_systems.ML
wenzelm@28592
     2
    Author:     Fabian Immler, TU Muenchen
blanchet@36371
     3
    Author:     Jasmin Blanchette, TU Muenchen
wenzelm@28592
     4
blanchet@36376
     5
Setup for supported ATPs.
wenzelm@28592
     6
*)
wenzelm@28592
     7
blanchet@36376
     8
signature ATP_SYSTEMS =
wenzelm@28592
     9
sig
blanchet@43448
    10
  type format = ATP_Problem.format
blanchet@43448
    11
  type formula_kind = ATP_Problem.formula_kind
blanchet@39731
    12
  type failure = ATP_Proof.failure
blanchet@38257
    13
blanchet@40240
    14
  type atp_config =
blanchet@43449
    15
    {exec : string * string,
blanchet@43449
    16
     required_execs : (string * string) list,
blanchet@43517
    17
     arguments :
blanchet@44344
    18
       Proof.context -> bool -> string -> Time.time
blanchet@44220
    19
       -> (unit -> (string * real) list) -> string,
blanchet@43449
    20
     proof_delims : (string * string) list,
blanchet@43449
    21
     known_failures : (failure * string) list,
blanchet@43580
    22
     conj_sym_kind : formula_kind,
blanchet@43580
    23
     prem_kind : formula_kind,
blanchet@43449
    24
     formats : format list,
blanchet@44344
    25
     best_slices :
blanchet@44431
    26
       Proof.context -> (real * (bool * (int * string * string))) list}
blanchet@38257
    27
blanchet@44428
    28
  val e_smartN : string
blanchet@44428
    29
  val e_autoN : string
blanchet@44428
    30
  val e_fun_weightN : string
blanchet@44428
    31
  val e_sym_offset_weightN : string
blanchet@43517
    32
  val e_weight_method : string Config.T
blanchet@43517
    33
  val e_default_fun_weight : real Config.T
blanchet@43517
    34
  val e_fun_weight_base : real Config.T
blanchet@43517
    35
  val e_fun_weight_span : real Config.T
blanchet@43517
    36
  val e_default_sym_offs_weight : real Config.T
blanchet@43517
    37
  val e_sym_offs_weight_base : real Config.T
blanchet@43517
    38
  val e_sym_offs_weight_span : real Config.T
blanchet@43590
    39
  val spass_force_sos : bool Config.T
blanchet@43590
    40
  val vampire_force_sos : bool Config.T
blanchet@40240
    41
  val eN : string
blanchet@40240
    42
  val spassN : string
blanchet@40240
    43
  val vampireN : string
blanchet@43803
    44
  val leo2N : string
blanchet@43803
    45
  val satallaxN : string
blanchet@44963
    46
  val e_sineN : string
blanchet@40240
    47
  val snarkN : string
blanchet@44963
    48
  val e_tofofN : string
blanchet@43779
    49
  val waldmeisterN : string
blanchet@42609
    50
  val z3_atpN : string
blanchet@40241
    51
  val remote_prefix : string
blanchet@42609
    52
  val remote_atp :
blanchet@42609
    53
    string -> string -> string list -> (string * string) list
blanchet@43580
    54
    -> (failure * string) list -> formula_kind -> formula_kind -> format list
blanchet@44431
    55
    -> (Proof.context -> int * string) -> string * atp_config
blanchet@40240
    56
  val add_atp : string * atp_config -> theory -> theory
blanchet@40240
    57
  val get_atp : theory -> string -> atp_config
blanchet@42591
    58
  val supported_atps : theory -> string list
blanchet@40240
    59
  val is_atp_installed : theory -> string -> bool
blanchet@35867
    60
  val refresh_systems_on_tptp : unit -> unit
blanchet@35867
    61
  val setup : theory -> theory
wenzelm@28592
    62
end;
wenzelm@28592
    63
blanchet@36376
    64
structure ATP_Systems : ATP_SYSTEMS =
wenzelm@28592
    65
struct
wenzelm@28596
    66
blanchet@43448
    67
open ATP_Problem
blanchet@39731
    68
open ATP_Proof
blanchet@39731
    69
blanchet@40240
    70
(* ATP configuration *)
blanchet@35826
    71
blanchet@40240
    72
type atp_config =
blanchet@43449
    73
  {exec : string * string,
blanchet@43449
    74
   required_execs : (string * string) list,
blanchet@43517
    75
   arguments :
blanchet@44344
    76
     Proof.context -> bool -> string -> Time.time
blanchet@44344
    77
     -> (unit -> (string * real) list) -> string,
blanchet@43449
    78
   proof_delims : (string * string) list,
blanchet@43449
    79
   known_failures : (failure * string) list,
blanchet@43580
    80
   conj_sym_kind : formula_kind,
blanchet@43580
    81
   prem_kind : formula_kind,
blanchet@43449
    82
   formats : format list,
blanchet@44344
    83
   best_slices :
blanchet@44431
    84
     Proof.context -> (real * (bool * (int * string * string))) list}
boehmes@32864
    85
blanchet@43588
    86
(* "best_slices" must be found empirically, taking a wholistic approach since
blanchet@43588
    87
   the ATPs are run in parallel. The "real" components give the faction of the
blanchet@44431
    88
   time available given to the slice and should add up to 1.0. The "bool"
blanchet@43588
    89
   component indicates whether the slice's strategy is complete; the "int", the
blanchet@44431
    90
   preferred number of facts to pass; the first "string", the preferred type
blanchet@44437
    91
   system (which should be sound or quasi-sound); the second "string", extra
blanchet@44437
    92
   information to the prover (e.g., SOS or no SOS).
blanchet@43588
    93
blanchet@43588
    94
   The last slice should be the most "normal" one, because it will get all the
blanchet@44431
    95
   time available if the other slices fail early and also because it is used if
blanchet@44431
    96
   slicing is disabled (e.g., by the minimizer). *)
blanchet@43581
    97
blanchet@38307
    98
val known_perl_failures =
blanchet@38340
    99
  [(CantConnect, "HTTP error"),
blanchet@38340
   100
   (NoPerl, "env: perl"),
blanchet@38311
   101
   (NoLibwwwPerl, "Can't locate HTTP")]
wenzelm@28596
   102
blanchet@40240
   103
(* named ATPs *)
blanchet@40240
   104
blanchet@40240
   105
val eN = "e"
blanchet@40240
   106
val spassN = "spass"
blanchet@40240
   107
val vampireN = "vampire"
blanchet@42611
   108
val z3_atpN = "z3_atp"
blanchet@43803
   109
val leo2N = "leo2"
blanchet@43803
   110
val satallaxN = "satallax"
blanchet@44963
   111
val e_sineN = "e_sine"
blanchet@40240
   112
val snarkN = "snark"
blanchet@44963
   113
val e_tofofN = "e_tofof"
blanchet@43779
   114
val waldmeisterN = "waldmeister"
blanchet@40241
   115
val remote_prefix = "remote_"
wenzelm@28596
   116
blanchet@38257
   117
structure Data = Theory_Data
blanchet@38257
   118
(
blanchet@40240
   119
  type T = (atp_config * stamp) Symtab.table
blanchet@38257
   120
  val empty = Symtab.empty
blanchet@38257
   121
  val extend = I
blanchet@38257
   122
  fun merge data : T = Symtab.merge (eq_snd op =) data
blanchet@38257
   123
    handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
blanchet@38257
   124
)
blanchet@36370
   125
blanchet@44852
   126
fun to_secs min time = Int.max (min, (Time.toMilliseconds time + 999) div 1000)
blanchet@36140
   127
blanchet@44344
   128
val sosN = "sos"
blanchet@44344
   129
val no_sosN = "no_sos"
blanchet@44344
   130
blanchet@39731
   131
blanchet@40240
   132
(* E *)
wenzelm@28596
   133
blanchet@36369
   134
val tstp_proof_delims =
blanchet@43803
   135
  [("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation"),
blanchet@43803
   136
   ("% SZS output start CNFRefutation", "% SZS output end CNFRefutation")]
blanchet@36369
   137
blanchet@44344
   138
val e_smartN = "smart"
blanchet@43517
   139
val e_autoN = "auto"
blanchet@43517
   140
val e_fun_weightN = "fun_weight"
blanchet@43517
   141
val e_sym_offset_weightN = "sym_offset_weight"
blanchet@42589
   142
blanchet@43517
   143
val e_weight_method =
blanchet@44344
   144
  Attrib.setup_config_string @{binding atp_e_weight_method} (K e_smartN)
blanchet@42641
   145
(* FUDGE *)
blanchet@43517
   146
val e_default_fun_weight =
blanchet@43517
   147
  Attrib.setup_config_real @{binding atp_e_default_fun_weight} (K 20.0)
blanchet@43517
   148
val e_fun_weight_base =
blanchet@43517
   149
  Attrib.setup_config_real @{binding atp_e_fun_weight_base} (K 0.0)
blanchet@43517
   150
val e_fun_weight_span =
blanchet@43517
   151
  Attrib.setup_config_real @{binding atp_e_fun_weight_span} (K 40.0)
blanchet@43517
   152
val e_default_sym_offs_weight =
blanchet@43517
   153
  Attrib.setup_config_real @{binding atp_e_default_sym_offs_weight} (K 1.0)
blanchet@43517
   154
val e_sym_offs_weight_base =
blanchet@43517
   155
  Attrib.setup_config_real @{binding atp_e_sym_offs_weight_base} (K ~20.0)
blanchet@43517
   156
val e_sym_offs_weight_span =
blanchet@43517
   157
  Attrib.setup_config_real @{binding atp_e_sym_offs_weight_span} (K 60.0)
blanchet@42589
   158
blanchet@43314
   159
fun e_weight_method_case method fw sow =
blanchet@43517
   160
  if method = e_fun_weightN then fw
blanchet@43517
   161
  else if method = e_sym_offset_weightN then sow
blanchet@44349
   162
  else raise Fail ("unexpected " ^ quote method)
blanchet@42589
   163
blanchet@43517
   164
fun scaled_e_weight ctxt method w =
blanchet@43517
   165
  w * Config.get ctxt
blanchet@43517
   166
          (e_weight_method_case method e_fun_weight_span e_sym_offs_weight_span)
blanchet@43517
   167
  + Config.get ctxt
blanchet@43517
   168
        (e_weight_method_case method e_fun_weight_base e_sym_offs_weight_base)
blanchet@42589
   169
  |> Real.ceil |> signed_string_of_int
blanchet@41561
   170
blanchet@43517
   171
fun e_weight_arguments ctxt method weights =
blanchet@43517
   172
  if method = e_autoN then
blanchet@42589
   173
    "-xAutoDev"
blanchet@42589
   174
  else
blanchet@44489
   175
    (* supplied by Stephan Schulz *)
blanchet@41562
   176
    "--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \
blanchet@41562
   177
    \--destructive-er-aggressive --destructive-er --presat-simplify \
blanchet@41562
   178
    \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \
blanchet@41562
   179
    \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \
blanchet@43314
   180
    \-H'(4*" ^ e_weight_method_case method "FunWeight" "SymOffsetWeight" ^
blanchet@42589
   181
    "(SimulateSOS, " ^
blanchet@43517
   182
    (e_weight_method_case method e_default_fun_weight e_default_sym_offs_weight
blanchet@43517
   183
     |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^
blanchet@41562
   184
    ",20,1.5,1.5,1" ^
blanchet@43517
   185
    (weights ()
blanchet@43517
   186
     |> map (fn (s, w) => "," ^ s ^ ":" ^ scaled_e_weight ctxt method w)
blanchet@43517
   187
     |> implode) ^
blanchet@41562
   188
    "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\
blanchet@41562
   189
    \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\
blanchet@41562
   190
    \FIFOWeight(PreferProcessed))'"
blanchet@41561
   191
blanchet@43314
   192
fun is_old_e_version () = (string_ord (getenv "E_VERSION", "1.2w") = LESS)
blanchet@43314
   193
blanchet@43517
   194
fun effective_e_weight_method ctxt =
blanchet@43517
   195
  if is_old_e_version () then e_autoN else Config.get ctxt e_weight_method
blanchet@43314
   196
blanchet@40240
   197
val e_config : atp_config =
blanchet@38338
   198
  {exec = ("E_HOME", "eproof"),
blanchet@38338
   199
   required_execs = [],
blanchet@44220
   200
   arguments =
blanchet@44429
   201
     fn ctxt => fn _ => fn method => fn timeout => fn weights =>
blanchet@44344
   202
        "--tstp-in --tstp-out -l5 " ^ e_weight_arguments ctxt method weights ^
blanchet@44852
   203
        " -tAutoDev --silent --cpu-limit=" ^ string_of_int (to_secs 2 timeout),
blanchet@43803
   204
   proof_delims = tstp_proof_delims,
blanchet@36265
   205
   known_failures =
blanchet@38229
   206
     [(Unprovable, "SZS status: CounterSatisfiable"),
blanchet@38229
   207
      (Unprovable, "SZS status CounterSatisfiable"),
blanchet@43713
   208
      (ProofMissing, "SZS status Theorem"),
blanchet@36370
   209
      (TimedOut, "Failure: Resource limit exceeded (time)"),
blanchet@36370
   210
      (TimedOut, "time limit exceeded"),
blanchet@44328
   211
      (OutOfResources, "# Cannot determine problem status"),
blanchet@36370
   212
      (OutOfResources, "SZS status: ResourceOut"),
blanchet@36370
   213
      (OutOfResources, "SZS status ResourceOut")],
blanchet@44327
   214
   conj_sym_kind = Hypothesis,
blanchet@43580
   215
   prem_kind = Conjecture,
blanchet@43778
   216
   formats = [FOF],
blanchet@43517
   217
   best_slices = fn ctxt =>
blanchet@44344
   218
     let val method = effective_e_weight_method ctxt in
blanchet@44345
   219
       (* FUDGE *)
blanchet@44344
   220
       if method = e_smartN then
blanchet@44431
   221
         [(0.333, (true, (500, "mangled_tags?", e_fun_weightN))),
blanchet@44860
   222
          (0.334, (true, (50, "mangled_guards?", e_fun_weightN))),
blanchet@44431
   223
          (0.333, (true, (1000, "mangled_tags?", e_sym_offset_weightN)))]
blanchet@44344
   224
       else
blanchet@44431
   225
         [(1.0, (true, (500, "mangled_tags?", method)))]
blanchet@44344
   226
     end}
blanchet@38698
   227
blanchet@40240
   228
val e = (eN, e_config)
wenzelm@28596
   229
wenzelm@28596
   230
blanchet@39731
   231
(* SPASS *)
blanchet@39731
   232
blanchet@43590
   233
val spass_force_sos =
blanchet@43590
   234
  Attrib.setup_config_bool @{binding atp_spass_force_sos} (K false)
blanchet@43590
   235
blanchet@36219
   236
(* The "-VarWeight=3" option helps the higher-order problems, probably by
blanchet@36219
   237
   counteracting the presence of "hAPP". *)
blanchet@40240
   238
val spass_config : atp_config =
blanchet@38338
   239
  {exec = ("ISABELLE_ATP", "scripts/spass"),
blanchet@39246
   240
   required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")],
blanchet@44431
   241
   arguments = fn _ => fn _ => fn sos => fn timeout => fn _ =>
blanchet@38199
   242
     ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
blanchet@44852
   243
      \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 1 timeout))
blanchet@44344
   244
     |> sos = sosN ? prefix "-SOS=1 ",
blanchet@36369
   245
   proof_delims = [("Here is a proof", "Formulae used in the proof")],
blanchet@36289
   246
   known_failures =
blanchet@38307
   247
     known_perl_failures @
blanchet@43891
   248
     [(GaveUp, "SPASS beiseite: Completion found"),
blanchet@36370
   249
      (TimedOut, "SPASS beiseite: Ran out of time"),
blanchet@36958
   250
      (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
blanchet@37388
   251
      (MalformedInput, "Undefined symbol"),
blanchet@37389
   252
      (MalformedInput, "Free Variable"),
blanchet@39490
   253
      (SpassTooOld, "tptp2dfg"),
blanchet@39490
   254
      (InternalError, "Please report this error")],
blanchet@44327
   255
   conj_sym_kind = Hypothesis,
blanchet@43580
   256
   prem_kind = Conjecture,
blanchet@43778
   257
   formats = [FOF],
blanchet@43590
   258
   best_slices = fn ctxt =>
blanchet@43588
   259
     (* FUDGE *)
blanchet@44431
   260
     [(0.333, (false, (150, "mangled_tags", sosN))),
blanchet@44431
   261
      (0.333, (false, (300, "poly_tags?", sosN))),
blanchet@44431
   262
      (0.334, (true, (50, "mangled_tags?", no_sosN)))]
blanchet@43590
   263
     |> (if Config.get ctxt spass_force_sos then hd #> apfst (K 1.0) #> single
blanchet@43590
   264
         else I)}
blanchet@38698
   265
blanchet@40240
   266
val spass = (spassN, spass_config)
wenzelm@28596
   267
blanchet@38698
   268
blanchet@37509
   269
(* Vampire *)
blanchet@37509
   270
blanchet@43590
   271
val vampire_force_sos =
blanchet@43590
   272
  Attrib.setup_config_bool @{binding atp_vampire_force_sos} (K false)
blanchet@43590
   273
blanchet@40240
   274
val vampire_config : atp_config =
blanchet@38338
   275
  {exec = ("VAMPIRE_HOME", "vampire"),
blanchet@38338
   276
   required_execs = [],
blanchet@44431
   277
   arguments = fn _ => fn _ => fn sos => fn timeout => fn _ =>
blanchet@44852
   278
     "--proof tptp --mode casc -t " ^ string_of_int (to_secs 1 timeout) ^
blanchet@41451
   279
     " --thanks \"Andrei and Krystof\" --input_file"
blanchet@44344
   280
     |> sos = sosN ? prefix "--sos on ",
blanchet@37509
   281
   proof_delims =
blanchet@37509
   282
     [("=========== Refutation ==========",
blanchet@37509
   283
       "======= End of refutation ======="),
blanchet@38279
   284
      ("% SZS output start Refutation", "% SZS output end Refutation"),
blanchet@38279
   285
      ("% SZS output start Proof", "% SZS output end Proof")],
blanchet@37509
   286
   known_failures =
blanchet@43891
   287
     [(GaveUp, "UNPROVABLE"),
blanchet@43891
   288
      (GaveUp, "CANNOT PROVE"),
blanchet@43891
   289
      (GaveUp, "SZS status GaveUp"),
blanchet@43751
   290
      (ProofIncomplete, "predicate_definition_introduction,[]"),
blanchet@38338
   291
      (TimedOut, "SZS status Timeout"),
blanchet@37509
   292
      (Unprovable, "Satisfiability detected"),
blanchet@38885
   293
      (Unprovable, "Termination reason: Satisfiable"),
blanchet@39490
   294
      (VampireTooOld, "not a valid option"),
blanchet@39490
   295
      (Interrupted, "Aborted by signal SIGINT")],
blanchet@44327
   296
   conj_sym_kind = Conjecture,
blanchet@43580
   297
   prem_kind = Conjecture,
blanchet@43778
   298
   formats = [FOF],
blanchet@43590
   299
   best_slices = fn ctxt =>
blanchet@43588
   300
     (* FUDGE *)
blanchet@44860
   301
     [(0.333, (false, (150, "poly_guards", sosN))),
blanchet@44860
   302
      (0.334, (true, (50, "mangled_guards?", no_sosN))),
blanchet@44431
   303
      (0.333, (false, (500, "mangled_tags?", sosN)))]
blanchet@43590
   304
     |> (if Config.get ctxt vampire_force_sos then hd #> apfst (K 1.0) #> single
blanchet@43590
   305
         else I)}
blanchet@38698
   306
blanchet@40240
   307
val vampire = (vampireN, vampire_config)
blanchet@37509
   308
blanchet@38698
   309
blanchet@42611
   310
(* Z3 with TPTP syntax *)
blanchet@42611
   311
blanchet@42611
   312
val z3_atp_config : atp_config =
blanchet@42611
   313
  {exec = ("Z3_HOME", "z3"),
blanchet@42611
   314
   required_execs = [],
blanchet@44220
   315
   arguments = fn _ => fn _ => fn _ => fn timeout => fn _ =>
blanchet@44852
   316
     "MBQI=true -p -t:" ^ string_of_int (to_secs 1 timeout),
blanchet@42611
   317
   proof_delims = [],
blanchet@42611
   318
   known_failures =
blanchet@42613
   319
     [(Unprovable, "\nsat"),
blanchet@43891
   320
      (GaveUp, "\nunknown"),
blanchet@43891
   321
      (GaveUp, "SZS status Satisfiable"),
blanchet@43314
   322
      (ProofMissing, "\nunsat"),
blanchet@43314
   323
      (ProofMissing, "SZS status Unsatisfiable")],
blanchet@43580
   324
   conj_sym_kind = Hypothesis,
blanchet@43580
   325
   prem_kind = Hypothesis,
blanchet@43778
   326
   formats = [FOF],
blanchet@43588
   327
   best_slices =
blanchet@43588
   328
     (* FUDGE (FIXME) *)
blanchet@44860
   329
     K [(0.5, (false, (250, "mangled_guards?", ""))),
blanchet@44860
   330
        (0.25, (false, (125, "mangled_guards?", ""))),
blanchet@44860
   331
        (0.125, (false, (62, "mangled_guards?", ""))),
blanchet@44860
   332
        (0.125, (false, (31, "mangled_guards?", "")))]}
blanchet@42611
   333
blanchet@42611
   334
val z3_atp = (z3_atpN, z3_atp_config)
blanchet@42611
   335
blanchet@42611
   336
blanchet@40240
   337
(* Remote ATP invocation via SystemOnTPTP *)
wenzelm@28596
   338
blanchet@38307
   339
val systems = Synchronized.var "atp_systems" ([] : string list)
immler@31828
   340
immler@31828
   341
fun get_systems () =
wenzelm@44721
   342
  case Isabelle_System.bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of
blanchet@39731
   343
    (output, 0) => split_lines output
blanchet@39731
   344
  | (output, _) =>
blanchet@39731
   345
    error (case extract_known_failure known_perl_failures output of
blanchet@42615
   346
             SOME failure => string_for_failure failure
blanchet@39731
   347
           | NONE => perhaps (try (unsuffix "\n")) output ^ ".")
immler@31828
   348
blanchet@43408
   349
fun find_system name [] systems =
blanchet@43408
   350
    find_first (String.isPrefix (name ^ "---")) systems
blanchet@38929
   351
  | find_system name (version :: versions) systems =
blanchet@38929
   352
    case find_first (String.isPrefix (name ^ "---" ^ version)) systems of
blanchet@38929
   353
      NONE => find_system name versions systems
blanchet@38929
   354
    | res => res
blanchet@38929
   355
blanchet@38929
   356
fun get_system name versions =
blanchet@38812
   357
  Synchronized.change_result systems
blanchet@38812
   358
      (fn systems => (if null systems then get_systems () else systems)
blanchet@43796
   359
                     |> `(`(find_system name versions)))
immler@31828
   360
blanchet@38929
   361
fun the_system name versions =
blanchet@38929
   362
  case get_system name versions of
blanchet@43796
   363
    (SOME sys, _) => sys
blanchet@43796
   364
  | (NONE, []) => error ("SystemOnTPTP is currently not available.")
blanchet@43796
   365
  | (NONE, syss) =>
blanchet@43796
   366
    error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n" ^
blanchet@43796
   367
           "(Available systems: " ^ commas_quote syss ^ ".)")
wenzelm@28596
   368
blanchet@41396
   369
val max_remote_secs = 240 (* give Geoff Sutcliffe's servers a break *)
blanchet@41396
   370
blanchet@38929
   371
fun remote_config system_name system_versions proof_delims known_failures
blanchet@43588
   372
                  conj_sym_kind prem_kind formats best_slice : atp_config =
blanchet@38338
   373
  {exec = ("ISABELLE_ATP", "scripts/remote_atp"),
blanchet@38338
   374
   required_execs = [],
blanchet@44220
   375
   arguments = fn _ => fn _ => fn _ => fn timeout => fn _ =>
blanchet@44852
   376
     "-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout))
blanchet@41396
   377
     ^ " -s " ^ the_system system_name system_versions,
blanchet@43803
   378
   proof_delims = union (op =) tstp_proof_delims proof_delims,
blanchet@43806
   379
   known_failures = known_failures @ known_perl_failures @
blanchet@43782
   380
     [(Unprovable, "says Satisfiable"),
blanchet@43840
   381
      (Unprovable, "says CounterSatisfiable"),
blanchet@43891
   382
      (GaveUp, "says Unknown"),
blanchet@43891
   383
      (GaveUp, "says GaveUp"),
blanchet@43806
   384
      (ProofMissing, "says Theorem"),
blanchet@43840
   385
      (ProofMissing, "says Unsatisfiable"),
blanchet@43794
   386
      (TimedOut, "says Timeout"),
blanchet@43794
   387
      (Inappropriate, "says Inappropriate")],
blanchet@43580
   388
   conj_sym_kind = conj_sym_kind,
blanchet@43580
   389
   prem_kind = prem_kind,
blanchet@43449
   390
   formats = formats,
blanchet@44344
   391
   best_slices = fn ctxt =>
blanchet@44493
   392
     let val (max_relevant, type_enc) = best_slice ctxt in
blanchet@44493
   393
       [(1.0, (false, (max_relevant, type_enc, "")))]
blanchet@44344
   394
     end}
blanchet@43314
   395
blanchet@44371
   396
fun remotify_config system_name system_versions best_slice
blanchet@44371
   397
        ({proof_delims, known_failures, conj_sym_kind, prem_kind, formats, ...}
blanchet@44371
   398
         : atp_config) : atp_config =
blanchet@38929
   399
  remote_config system_name system_versions proof_delims known_failures
blanchet@44371
   400
                conj_sym_kind prem_kind formats best_slice
blanchet@38257
   401
blanchet@40240
   402
fun remote_atp name system_name system_versions proof_delims known_failures
blanchet@43588
   403
        conj_sym_kind prem_kind formats best_slice =
blanchet@40241
   404
  (remote_prefix ^ name,
blanchet@38929
   405
   remote_config system_name system_versions proof_delims known_failures
blanchet@43588
   406
                 conj_sym_kind prem_kind formats best_slice)
blanchet@44371
   407
fun remotify_atp (name, config) system_name system_versions best_slice =
blanchet@44371
   408
  (remote_prefix ^ name,
blanchet@44371
   409
   remotify_config system_name system_versions best_slice config)
wenzelm@28592
   410
blanchet@44371
   411
val remote_e =
blanchet@44371
   412
  remotify_atp e "EP" ["1.0", "1.1", "1.2"]
blanchet@44431
   413
               (K (750, "mangled_tags?") (* FUDGE *))
blanchet@44371
   414
val remote_vampire =
blanchet@44371
   415
  remotify_atp vampire "Vampire" ["0.6", "9.0", "1.0"]
blanchet@44860
   416
               (K (200, "mangled_guards?") (* FUDGE *))
blanchet@44371
   417
val remote_z3_atp =
blanchet@44860
   418
  remotify_atp z3_atp "Z3" ["2.18"] (K (250, "mangled_guards?") (* FUDGE *))
blanchet@43803
   419
val remote_leo2 =
blanchet@44967
   420
  remote_atp leo2N "LEO-II" ["1.2.6"] [] [] Axiom Hypothesis [THF, FOF]
blanchet@44491
   421
             (K (100, "simple_higher") (* FUDGE *))
blanchet@43803
   422
val remote_satallax =
blanchet@43803
   423
  remote_atp satallaxN "Satallax" ["2.0"] [] [] Axiom Hypothesis [THF]
blanchet@44491
   424
             (K (64, "simple_higher") (* FUDGE *))
blanchet@44963
   425
val remote_e_sine =
blanchet@44963
   426
  remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Axiom
blanchet@44860
   427
             Conjecture [FOF] (K (500, "mangled_guards?") (* FUDGE *))
blanchet@42611
   428
val remote_snark =
blanchet@43780
   429
  remote_atp snarkN "SNARK" ["20080805r029", "20080805r024"]
blanchet@43803
   430
             [("refutation.", "end_refutation.")] [] Hypothesis Hypothesis
blanchet@44431
   431
             [TFF, FOF] (K (100, "simple") (* FUDGE *))
blanchet@44963
   432
val remote_e_tofof =
blanchet@44963
   433
  remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config)
blanchet@44431
   434
             Axiom Hypothesis [TFF] (K (150, "simple") (* FUDGE *))
blanchet@43779
   435
val remote_waldmeister =
blanchet@43779
   436
  remote_atp waldmeisterN "Waldmeister" ["710"]
blanchet@43785
   437
             [("#START OF PROOF", "Proved Goals:")]
blanchet@43794
   438
             [(OutOfResources, "Too many function symbols"),
blanchet@43794
   439
              (Crashed, "Unrecoverable Segmentation Fault")]
blanchet@43794
   440
             Hypothesis Hypothesis [CNF_UEQ]
blanchet@44431
   441
             (K (50, "mangled_tags?") (* FUDGE *))
blanchet@38698
   442
blanchet@38698
   443
(* Setup *)
blanchet@38698
   444
blanchet@40240
   445
fun add_atp (name, config) thy =
blanchet@40240
   446
  Data.map (Symtab.update_new (name, (config, stamp ()))) thy
blanchet@40240
   447
  handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
blanchet@36371
   448
blanchet@40240
   449
fun get_atp thy name =
blanchet@40240
   450
  the (Symtab.lookup (Data.get thy) name) |> fst
blanchet@40240
   451
  handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
blanchet@36371
   452
blanchet@42591
   453
val supported_atps = Symtab.keys o Data.get
blanchet@40240
   454
blanchet@40240
   455
fun is_atp_installed thy name =
blanchet@40240
   456
  let val {exec, required_execs, ...} = get_atp thy name in
blanchet@40240
   457
    forall (curry (op <>) "" o getenv o fst) (exec :: required_execs)
blanchet@40240
   458
  end
blanchet@40240
   459
blanchet@40240
   460
fun refresh_systems_on_tptp () =
blanchet@40240
   461
  Synchronized.change systems (fn _ => get_systems ())
blanchet@40240
   462
blanchet@43803
   463
val atps =
blanchet@43803
   464
  [e, spass, vampire, z3_atp, remote_e, remote_vampire, remote_z3_atp,
blanchet@44963
   465
   remote_leo2, remote_satallax, remote_e_sine, remote_snark, remote_e_tofof,
blanchet@43804
   466
   remote_waldmeister]
blanchet@40240
   467
val setup = fold add_atp atps
blanchet@35867
   468
wenzelm@28592
   469
end;