src/HOL/Tools/Sledgehammer/sledgehammer_atp_systems.ML
author desharna
Wed, 22 Sep 2021 12:41:40 +0200
changeset 74352 fb8ce6090437
parent 74351 d8dbe7525ff1
child 74389 c1583aa3d861
permissions -rw-r--r--
removed checks for non-commercial usage of Vampire as it is now under BSD licence
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38047
9033c03cc214 consequence of directory renaming
blanchet
parents: 38046
diff changeset
     1
(*  Title:      HOL/Tools/ATP/atp_systems.ML
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
     2
    Author:     Fabian Immler, TU Muenchen
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
     4
36376
e83d52a52449 renamed module "ATP_Wrapper" to "ATP_Systems"
blanchet
parents: 36371
diff changeset
     5
Setup for supported ATPs.
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
     6
*)
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
     7
72400
abfeed05c323 tune filename
desharna
parents: 72174
diff changeset
     8
signature SLEDGEHAMMER_ATP_SYSTEMS =
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
     9
sig
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
    10
  type term_order = ATP_Problem.term_order
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
    11
  type atp_format = ATP_Problem.atp_format
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53515
diff changeset
    12
  type atp_formula_role = ATP_Problem.atp_formula_role
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53515
diff changeset
    13
  type atp_failure = ATP_Proof.atp_failure
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    14
51011
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    15
  type slice_spec = (int * string) * atp_format * string * string * bool
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    16
  type atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
    17
    {exec : string list * string list,
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
    18
     arguments : Proof.context -> bool -> string -> Time.time -> Path.T ->
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
    19
       term_order * (unit -> (string * int) list) * (unit -> (string * real) list) -> string list,
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    20
     proof_delims : (string * string) list,
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53515
diff changeset
    21
     known_failures : (atp_failure * string) list,
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53515
diff changeset
    22
     prem_role : atp_formula_role,
48716
1d2a12bb0640 stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents: 48715
diff changeset
    23
     best_slices : Proof.context -> (real * (slice_spec * string)) list,
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    24
     best_max_mono_iters : int,
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    25
     best_max_new_mono_instances : int}
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    26
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    27
  val default_max_mono_iters : int
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    28
  val default_max_new_mono_instances : int
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
    29
  val force_sos : bool Config.T
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
    30
  val term_order : string Config.T
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43529
diff changeset
    31
  val e_smartN : string
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43529
diff changeset
    32
  val e_autoN : string
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43529
diff changeset
    33
  val e_fun_weightN : string
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43529
diff changeset
    34
  val e_sym_offset_weightN : string
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
    35
  val e_selection_heuristic : string Config.T
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    36
  val e_default_fun_weight : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    37
  val e_fun_weight_base : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    38
  val e_fun_weight_span : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    39
  val e_default_sym_offs_weight : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    40
  val e_sym_offs_weight_base : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    41
  val e_sym_offs_weight_span : real Config.T
50950
a145ee10371b make SPASS more configurable, for experiments
blanchet
parents: 50927
diff changeset
    42
  val spass_H1SOS : string
a145ee10371b make SPASS more configurable, for experiments
blanchet
parents: 50927
diff changeset
    43
  val spass_H2 : string
a145ee10371b make SPASS more configurable, for experiments
blanchet
parents: 50927
diff changeset
    44
  val spass_H2LR0LT0 : string
a145ee10371b make SPASS more configurable, for experiments
blanchet
parents: 50927
diff changeset
    45
  val spass_H2NuVS0 : string
a145ee10371b make SPASS more configurable, for experiments
blanchet
parents: 50927
diff changeset
    46
  val spass_H2NuVS0Red2 : string
a145ee10371b make SPASS more configurable, for experiments
blanchet
parents: 50927
diff changeset
    47
  val spass_H2SOS : string
73435
1cc848548f21 invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents: 73432
diff changeset
    48
  val isabelle_scala_function: string list * string list
57671
dc5e1b1db9ba avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
blanchet
parents: 57636
diff changeset
    49
  val remote_atp : string -> string -> string list -> (string * string) list ->
dc5e1b1db9ba avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
blanchet
parents: 57636
diff changeset
    50
    (atp_failure * string) list -> atp_formula_role -> (Proof.context -> slice_spec * string) ->
dc5e1b1db9ba avoid 'eproof' and 'eproof_ram' scripts if possible (i.e. if 'eprover' can produce reasonable enough proofs for one-liner reconstruction)
blanchet
parents: 57636
diff changeset
    51
    string * (unit -> atp_config)
47606
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47506
diff changeset
    52
  val add_atp : string * (unit -> atp_config) -> theory -> theory
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47506
diff changeset
    53
  val get_atp : theory -> string -> (unit -> atp_config)
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41725
diff changeset
    54
  val supported_atps : theory -> string list
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    55
  val is_atp_installed : theory -> string -> bool
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
    56
  val refresh_systems_on_tptp : unit -> unit
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
    57
  val effective_term_order : Proof.context -> string -> term_order
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    58
end;
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    59
72400
abfeed05c323 tune filename
desharna
parents: 72174
diff changeset
    60
structure Sledgehammer_ATP_Systems : SLEDGEHAMMER_ATP_SYSTEMS =
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    61
struct
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
    62
42577
78414ec6fa4e made the format (TFF or FOF) of the TPTP problem a global argument of the problem again and have the ATPs report which formats they support
blanchet
parents: 42571
diff changeset
    63
open ATP_Problem
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
    64
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45876
diff changeset
    65
open ATP_Problem_Generate
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
    66
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
    67
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    68
(* ATP configuration *)
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
    69
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    70
val default_max_mono_iters = 3 (* FUDGE *)
53515
f5b678b155f6 adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents: 53225
diff changeset
    71
val default_max_new_mono_instances = 100 (* FUDGE *)
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    72
51011
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    73
type slice_spec = (int * string) * atp_format * string * string * bool
46409
d4754183ccce made option available to users (mostly for experiments)
blanchet
parents: 46407
diff changeset
    74
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    75
type atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
    76
  {exec : string list * string list,
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
    77
   arguments : Proof.context -> bool -> string -> Time.time -> Path.T ->
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
    78
     term_order * (unit -> (string * int) list) * (unit -> (string * real) list) -> string list,
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    79
   proof_delims : (string * string) list,
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53515
diff changeset
    80
   known_failures : (atp_failure * string) list,
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53515
diff changeset
    81
   prem_role : atp_formula_role,
48716
1d2a12bb0640 stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents: 48715
diff changeset
    82
   best_slices : Proof.context -> (real * (slice_spec * string)) list,
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    83
   best_max_mono_iters : int,
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
    84
   best_max_new_mono_instances : int}
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
    85
72401
2783779b7dd3 removed obsolete unmaintained experimental prover Pirate
blanchet
parents: 72400
diff changeset
    86
(* "best_slices" must be found empirically, taking a holistic approach since the
2783779b7dd3 removed obsolete unmaintained experimental prover Pirate
blanchet
parents: 72400
diff changeset
    87
   ATPs are run in parallel. Each slice has the format
51011
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    88
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    89
     (time_frac, ((max_facts, fact_filter), format, type_enc,
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    90
                  lam_trans, uncurried_aliases), extra)
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    91
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    92
   where
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    93
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    94
     time_frac = faction of the time available given to the slice (which should
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    95
       add up to 1.0)
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    96
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
    97
     extra = extra information to the prover (e.g., SOS or no SOS).
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
    98
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
    99
   The last slice should be the most "normal" one, because it will get all the
43569
b342cd125533 removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents: 43567
diff changeset
   100
   time available if the other slices fail early and also because it is used if
b342cd125533 removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents: 43567
diff changeset
   101
   slicing is disabled (e.g., by the minimizer). *)
42710
84fcce345b5d further improved type system setup based on Judgment Days
blanchet
parents: 42709
diff changeset
   102
51011
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
   103
val mepoN = "mepo"
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
   104
val mashN = "mash"
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
   105
val meshN = "mesh"
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
   106
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   107
val tstp_proof_delims =
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   108
  [("% SZS output start CNFRefutation", "% SZS output end CNFRefutation"),
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   109
   ("% SZS output start Refutation", "% SZS output end Refutation"),
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   110
   ("% SZS output start Proof", "% SZS output end Proof")]
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   111
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   112
fun known_szs_failures wrap =
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   113
  [(Unprovable, wrap "CounterSatisfiable"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   114
   (Unprovable, wrap "Satisfiable"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   115
   (GaveUp, wrap "GaveUp"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   116
   (GaveUp, wrap "Unknown"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   117
   (GaveUp, wrap "Incomplete"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   118
   (ProofMissing, wrap "Theorem"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   119
   (ProofMissing, wrap "Unsatisfiable"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   120
   (TimedOut, wrap "Timeout"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   121
   (Inappropriate, wrap "Inappropriate"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   122
   (OutOfResources, wrap "ResourceOut"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   123
   (OutOfResources, wrap "MemoryOut"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   124
   (Interrupted, wrap "Forced"),
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   125
   (Interrupted, wrap "User")]
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   126
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   127
val known_szs_status_failures = known_szs_failures (prefix "SZS status ")
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   128
val known_says_failures = known_szs_failures (prefix " says ")
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   129
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   130
structure Data = Theory_Data
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   131
(
47606
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47506
diff changeset
   132
  type T = ((unit -> atp_config) * stamp) Symtab.table
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   133
  val empty = Symtab.empty
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   134
  val extend = I
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   135
  fun merge data : T =
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   136
    Symtab.merge (eq_snd (op =)) data
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63116
diff changeset
   137
    handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name)
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   138
)
38017
3ad3e3ca2451 move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents: 38015
diff changeset
   139
43981
404ae49ce29f give E at least two seconds -- anything else risks causing too early timeouts in the minimizer, because of too conservative time computations in E and eproof scripts
blanchet
parents: 43850
diff changeset
   140
fun to_secs min time = Int.max (min, (Time.toMilliseconds time + 999) div 1000)
36142
f5e15e9aae10 make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents: 36064
diff changeset
   141
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   142
val sosN = "sos"
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   143
val no_sosN = "no_sos"
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   144
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   145
val force_sos = Attrib.setup_config_bool \<^binding>\<open>atp_force_sos\<close> (K false)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   146
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
   147
val smartN = "smart"
47073
c73f7b0c7ebc generate weights and precedences for predicates as well
blanchet
parents: 47055
diff changeset
   148
(* val kboN = "kbo" *)
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
   149
val lpoN = "lpo"
47034
77da780ddd6b implement term order attribute (for experiments)
blanchet
parents: 47033
diff changeset
   150
val xweightsN = "_weights"
77da780ddd6b implement term order attribute (for experiments)
blanchet
parents: 47033
diff changeset
   151
val xprecN = "_prec"
77da780ddd6b implement term order attribute (for experiments)
blanchet
parents: 47033
diff changeset
   152
val xsimpN = "_simp" (* SPASS-specific *)
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
   153
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   154
(* Possible values for "atp_term_order":
47049
72bd3311ecba added term_order option to Mirabelle
blanchet
parents: 47039
diff changeset
   155
   "smart", "(kbo|lpo)(_weights)?(_prec|_simp)?" *)
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
   156
val term_order =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   157
  Attrib.setup_config_string \<^binding>\<open>atp_term_order\<close> (K smartN)
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
   158
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   159
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   160
(* agsyHOL *)
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   161
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   162
val agsyhol_config : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   163
  {exec = (["AGSYHOL_HOME"], ["agsyHOL"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   164
   arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   165
     ["--proof --time-out " ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem],
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   166
   proof_delims = tstp_proof_delims,
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   167
   known_failures = known_szs_status_failures,
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   168
   prem_role = Hypothesis,
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   169
   best_slices =
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   170
     (* FUDGE *)
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   171
     K [(1.0, (((60, ""), THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))],
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   172
   best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *),
53515
f5b678b155f6 adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents: 53225
diff changeset
   173
   best_max_new_mono_instances = default_max_new_mono_instances}
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   174
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   175
val agsyhol = (agsyholN, fn () => agsyhol_config)
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   176
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   177
46643
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   178
(* Alt-Ergo *)
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   179
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   180
val alt_ergo_config : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   181
  {exec = (["WHY3_HOME"], ["why3"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   182
   arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   183
     ["--format tptp --prover 'Alt-Ergo,0.95.2,' --timelimit " ^ string_of_int (to_secs 1 timeout) ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   184
      " " ^ File.bash_path problem],
46643
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   185
   proof_delims = [],
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   186
   known_failures =
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   187
     [(ProofMissing, ": Valid"),
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   188
      (TimedOut, ": Timeout"),
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   189
      (GaveUp, ": Unknown")],
47976
6b13451135a9 tuned names
blanchet
parents: 47974
diff changeset
   190
   prem_role = Hypothesis,
46643
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   191
   best_slices = fn _ =>
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   192
     (* FUDGE *)
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   193
     [(1.0, (((100, ""), TFF (Without_FOOL, Polymorphic), "poly_native", liftingN, false), ""))],
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   194
   best_max_mono_iters = default_max_mono_iters,
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   195
   best_max_new_mono_instances = default_max_new_mono_instances}
46643
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   196
47646
9460f3f22365 tried to make SML/NJ happy
blanchet
parents: 47606
diff changeset
   197
val alt_ergo = (alt_ergoN, fn () => alt_ergo_config)
46643
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   198
a88bccd2b567 added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents: 46481
diff changeset
   199
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   200
(* E *)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   201
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   202
val e_smartN = "smart"
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   203
val e_autoN = "auto"
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   204
val e_fun_weightN = "fun_weight"
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   205
val e_sym_offset_weightN = "sym_offset_weight"
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   206
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47031
diff changeset
   207
val e_selection_heuristic =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   208
  Attrib.setup_config_string \<^binding>\<open>atp_e_selection_heuristic\<close> (K e_smartN)
41770
a710e96583d5 adjust fudge factors
blanchet
parents: 41769
diff changeset
   209
(* FUDGE *)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   210
val e_default_fun_weight =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   211
  Attrib.setup_config_real \<^binding>\<open>atp_e_default_fun_weight\<close> (K 20.0)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   212
val e_fun_weight_base =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   213
  Attrib.setup_config_real \<^binding>\<open>atp_e_fun_weight_base\<close> (K 0.0)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   214
val e_fun_weight_span =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   215
  Attrib.setup_config_real \<^binding>\<open>atp_e_fun_weight_span\<close> (K 40.0)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   216
val e_default_sym_offs_weight =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   217
  Attrib.setup_config_real \<^binding>\<open>atp_e_default_sym_offs_weight\<close> (K 1.0)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   218
val e_sym_offs_weight_base =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   219
  Attrib.setup_config_real \<^binding>\<open>atp_e_sym_offs_weight_base\<close> (K ~20.0)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   220
val e_sym_offs_weight_span =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68563
diff changeset
   221
  Attrib.setup_config_real \<^binding>\<open>atp_e_sym_offs_weight_span\<close> (K 60.0)
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   222
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   223
fun e_selection_heuristic_case heuristic fw sow =
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   224
  if heuristic = e_fun_weightN then fw
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   225
  else if heuristic = e_sym_offset_weightN then sow
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   226
  else raise Fail ("unexpected " ^ quote heuristic)
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   227
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   228
fun scaled_e_selection_weight ctxt heuristic w =
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   229
  w * Config.get ctxt (e_selection_heuristic_case heuristic
47029
72802e2edda4 renamed E weight attribute
blanchet
parents: 46643
diff changeset
   230
                           e_fun_weight_span e_sym_offs_weight_span)
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   231
  + Config.get ctxt (e_selection_heuristic_case heuristic
47029
72802e2edda4 renamed E weight attribute
blanchet
parents: 46643
diff changeset
   232
                         e_fun_weight_base e_sym_offs_weight_base)
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   233
  |> Real.ceil |> signed_string_of_int
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41269
diff changeset
   234
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   235
fun e_selection_weight_arguments ctxt heuristic sel_weights =
51631
8d60dfb41d19 robustness w.r.t. unknown arguments
blanchet
parents: 51467
diff changeset
   236
  if heuristic = e_fun_weightN orelse heuristic = e_sym_offset_weightN then
43622
blanchet
parents: 43575
diff changeset
   237
    (* supplied by Stephan Schulz *)
41314
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   238
    "--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   239
    \--destructive-er-aggressive --destructive-er --presat-simplify \
47505
e33d957ae2bf avoid option introduced in E 1.2 when invoking older versions of E
blanchet
parents: 47499
diff changeset
   240
    \--prefer-initial-clauses -winvfreqrank -c1 -Ginvfreqconjmax -F1 \
e33d957ae2bf avoid option introduced in E 1.2 when invoking older versions of E
blanchet
parents: 47499
diff changeset
   241
    \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred -H'(4*" ^
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   242
    e_selection_heuristic_case heuristic "FunWeight" "SymOffsetWeight" ^
48376
416e4123baf3 use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents: 48232
diff changeset
   243
    "(SimulateSOS," ^
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   244
    (e_selection_heuristic_case heuristic
47029
72802e2edda4 renamed E weight attribute
blanchet
parents: 46643
diff changeset
   245
         e_default_fun_weight e_default_sym_offs_weight
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   246
     |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^
41314
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   247
    ",20,1.5,1.5,1" ^
47030
7e80e14247fc internal renamings
blanchet
parents: 47029
diff changeset
   248
    (sel_weights ()
47029
72802e2edda4 renamed E weight attribute
blanchet
parents: 46643
diff changeset
   249
     |> map (fn (s, w) => "," ^ s ^ ":" ^
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   250
                          scaled_e_selection_weight ctxt heuristic w)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   251
     |> implode) ^
41314
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   252
    "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   253
    \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\
57672
blanchet
parents: 57671
diff changeset
   254
    \FIFOWeight(PreferProcessed))' "
51631
8d60dfb41d19 robustness w.r.t. unknown arguments
blanchet
parents: 51467
diff changeset
   255
  else
57672
blanchet
parents: 57671
diff changeset
   256
    "-xAuto "
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41269
diff changeset
   257
70939
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   258
val e_ord_weights = map (fn (s, w) => s ^ ":" ^ string_of_int w) #> space_implode ","
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   259
fun e_ord_precedence [_] = ""
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   260
  | e_ord_precedence info = info |> map fst |> space_implode "<"
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   261
47039
1b36a05a070d added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
blanchet
parents: 47038
diff changeset
   262
fun e_term_order_info_arguments false false _ = ""
1b36a05a070d added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
blanchet
parents: 47038
diff changeset
   263
  | e_term_order_info_arguments gen_weights gen_prec ord_info =
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   264
    let val ord_info = ord_info () in
57672
blanchet
parents: 57671
diff changeset
   265
      (if gen_weights then "--order-weights='" ^ e_ord_weights ord_info ^ "' " else "") ^
blanchet
parents: 57671
diff changeset
   266
      (if gen_prec then "--precedence='" ^ e_ord_precedence ord_info ^ "' " else "")
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   267
    end
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   268
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   269
val e_config : atp_config =
73973
f0d231ead660 added alternative E binary name
blanchet
parents: 73970
diff changeset
   270
  {exec = (["E_HOME"], ["eprover-ho", "eprover"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   271
   arguments = fn ctxt => fn _ => fn heuristic => fn timeout => fn problem =>
57008
10f68b83b474 use E 1.8's auto scheduler option
blanchet
parents: 56848
diff changeset
   272
     fn ({is_lpo, gen_weights, gen_prec, ...}, ord_info, sel_weights) =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   273
       ["--auto-schedule --tstp-in --tstp-out --silent " ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   274
        e_selection_weight_arguments ctxt heuristic sel_weights ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   275
        e_term_order_info_arguments gen_weights gen_prec ord_info ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   276
        "--term-ordering=" ^ (if is_lpo then "LPO4" else "KBO6") ^ " " ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   277
        "--cpu-limit=" ^ string_of_int (to_secs 2 timeout) ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   278
        " --proof-object=1 " ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   279
        File.bash_path problem],
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   280
   proof_delims =
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   281
     [("# SZS output start CNFRefutation", "# SZS output end CNFRefutation")] @
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   282
     tstp_proof_delims,
36265
41c9e755e552 distinguish between the different ATP errors in the user interface;
blanchet
parents: 36264
diff changeset
   283
   known_failures =
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   284
     [(TimedOut, "Failure: Resource limit exceeded (time)"),
47972
96c9b8381909 better handling of incomplete TSTP proofs
blanchet
parents: 47962
diff changeset
   285
      (TimedOut, "time limit exceeded")] @
96c9b8381909 better handling of incomplete TSTP proofs
blanchet
parents: 47962
diff changeset
   286
     known_szs_status_failures,
47976
6b13451135a9 tuned names
blanchet
parents: 47974
diff changeset
   287
   prem_role = Conjecture,
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   288
   best_slices = fn ctxt =>
70939
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   289
     let
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   290
       val heuristic = Config.get ctxt e_selection_heuristic
74047
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   291
       val (format, enc, main_lam_trans) =
73974
6a0e1c14a8c2 wait for E 2.7 before using 'ite' in HO mode
blanchet
parents: 73973
diff changeset
   292
         if string_ord (getenv "E_VERSION", "2.7") <> LESS then
74047
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   293
           (THF (With_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher_fool", keep_lamsN)
73974
6a0e1c14a8c2 wait for E 2.7 before using 'ite' in HO mode
blanchet
parents: 73973
diff changeset
   294
         else if string_ord (getenv "E_VERSION", "2.6") <> LESS then
74047
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   295
           (THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN)
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   296
         else
74047
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   297
           (THF (Without_FOOL, Monomorphic, THF_Lambda_Free), "mono_native_higher", combsN)
70939
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   298
     in
43474
423cd1ecf714 optimized E's time slicing, based on latest exhaustive Judgment Day results
blanchet
parents: 43473
diff changeset
   299
       (* FUDGE *)
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 47034
diff changeset
   300
       if heuristic = e_smartN then
74047
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   301
         [(0.15, (((128, meshN), format, enc, main_lam_trans, false), e_fun_weightN)),
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   302
          (0.15, (((128, mashN), format, enc, main_lam_trans, false), e_sym_offset_weightN)),
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   303
          (0.15, (((91, mepoN), format, enc, main_lam_trans, false), e_autoN)),
a2b470e315ee tuned E's lambda encoding
blanchet
parents: 74046
diff changeset
   304
          (0.15, (((1000, meshN), format, "poly_guards??", main_lam_trans, false), e_sym_offset_weightN)),
70939
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   305
          (0.15, (((256, mepoN), format, enc, liftingN, false), e_fun_weightN)),
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   306
          (0.25, (((64, mashN), format, enc, combsN, false), e_fun_weightN))]
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   307
       else
70939
3218999b3715 folded experimental Ehoh into E now that E 2.3 has been released
blanchet
parents: 70938
diff changeset
   308
         [(1.0, (((500, ""), format, enc, combsN, false), heuristic))]
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   309
     end,
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   310
   best_max_mono_iters = default_max_mono_iters,
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   311
   best_max_new_mono_instances = default_max_new_mono_instances}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   312
47646
9460f3f22365 tried to make SML/NJ happy
blanchet
parents: 47606
diff changeset
   313
val e = (eN, fn () => e_config)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   314
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   315
48700
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   316
(* iProver *)
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   317
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   318
val iprover_config : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   319
  {exec = (["IPROVER_HOME"], ["iproveropt", "iprover"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   320
   arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => fn _ =>
74350
398b7bb9ebdd used Vampire 4.5.1 in Sledgehammer
desharna
parents: 74311
diff changeset
   321
     ["--clausifier \"$VAMPIRE_HOME\"/vampire " ^
74046
462d652ad910 use Vampire's clausifier with iProver, now that E's is no longer supported
blanchet
parents: 74005
diff changeset
   322
      "--clausifier_options \"--mode clausify\" " ^
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   323
      "--time_out_real " ^ string_of_real (Time.toReal timeout) ^ " " ^ File.bash_path problem],
48700
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   324
   proof_delims = tstp_proof_delims,
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   325
   known_failures =
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   326
     [(ProofIncomplete, "% SZS output start CNFRefutation")] @
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   327
     known_szs_status_failures,
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   328
   prem_role = Hypothesis,
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   329
   best_slices =
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   330
     (* FUDGE *)
51011
62b992e53eb8 store fact filter along with ATP slice
blanchet
parents: 50950
diff changeset
   331
     K [(1.0, (((150, ""), FOF, "mono_guards??", liftingN, false), ""))],
48700
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   332
   best_max_mono_iters = default_max_mono_iters,
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   333
   best_max_new_mono_instances = default_max_new_mono_instances}
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   334
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   335
val iprover = (iproverN, fn () => iprover_config)
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   336
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   337
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   338
(* LEO-II *)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   339
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   340
val leo2_config : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   341
  {exec = (["LEO2_HOME"], ["leo.opt", "leo"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   342
   arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   343
     ["--foatp e --atp e=\"$E_HOME\"/eprover \
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   344
      \--atp epclextract=\"$E_HOME\"/epclextract \
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   345
      \--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   346
      (if full_proofs then "--notReplLeibnizEQ --notReplAndrewsEQ --notUseExtCnfCmbd " else "") ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   347
      File.bash_path problem],
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   348
   proof_delims = tstp_proof_delims,
45207
1d13334628a9 one more LEO-II failure
blanchet
parents: 45203
diff changeset
   349
   known_failures =
47974
08d2dcc2dab9 improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents: 47972
diff changeset
   350
     [(TimedOut, "CPU time limit exceeded, terminating"),
47972
96c9b8381909 better handling of incomplete TSTP proofs
blanchet
parents: 47962
diff changeset
   351
      (GaveUp, "No.of.Axioms")] @
96c9b8381909 better handling of incomplete TSTP proofs
blanchet
parents: 47962
diff changeset
   352
     known_szs_status_failures,
47976
6b13451135a9 tuned names
blanchet
parents: 47974
diff changeset
   353
   prem_role = Hypothesis,
47914
94f37848b7c9 LEO-II's "--sos" option confusingly disables rather than enables SOS, and SOS seems to be ignored anyway; also, pass a number of facts that's more appropriate for each prover
blanchet
parents: 47912
diff changeset
   354
   best_slices =
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   355
     (* FUDGE *)
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   356
     K [(1.0, (((40, ""), THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))],
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   357
   best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *),
53515
f5b678b155f6 adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents: 53225
diff changeset
   358
   best_max_new_mono_instances = default_max_new_mono_instances}
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   359
47646
9460f3f22365 tried to make SML/NJ happy
blanchet
parents: 47606
diff changeset
   360
val leo2 = (leo2N, fn () => leo2_config)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   361
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   362
67021
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   363
(* Leo-III *)
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   364
69717
eb74ff534b27 tune ATP settings
blanchet
parents: 69593
diff changeset
   365
(* Include choice? Disabled now since it's disabled for Satallax as well. *)
67021
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   366
val leo3_config : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   367
  {exec = (["LEO3_HOME"], ["leo3"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   368
   arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   369
     [File.bash_path problem ^ " " ^ "--atp cvc=\"$CVC4_SOLVER\" --atp e=\"$E_HOME\"/eprover \
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   370
      \-p -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   371
      (if full_proofs then "--nleq --naeq " else "")],
67021
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   372
   proof_delims = tstp_proof_delims,
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   373
   known_failures = known_szs_status_failures,
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   374
   prem_role = Hypothesis,
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   375
   best_slices =
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   376
     (* FUDGE *)
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   377
     K [(1.0, (((150, ""), THF (Without_FOOL, Polymorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))],
67021
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   378
   best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *),
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   379
   best_max_new_mono_instances = default_max_new_mono_instances}
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   380
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   381
val leo3 = (leo3N, fn () => leo3_config)
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   382
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   383
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   384
(* Satallax *)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   385
52097
f353fe3c2b92 disabled choice in Satallax
blanchet
parents: 52095
diff changeset
   386
(* Choice is disabled until there is proper reconstruction for it. *)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   387
val satallax_config : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   388
  {exec = (["SATALLAX_HOME"], ["satallax.opt", "satallax"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   389
   arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   390
     [(case getenv "E_HOME" of
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   391
        "" => ""
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   392
      | home => "-E " ^ home ^ "/eprover ") ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   393
      "-p tstp -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem],
45162
170dffc6df75 parse Satallax unsat cores
blanchet
parents: 44786
diff changeset
   394
   proof_delims =
57707
0242e9578828 imported patch satallax_proof_support_Sledgehammer
fleury
parents: 57672
diff changeset
   395
     [("% SZS output start Proof", "% SZS output end Proof")],
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   396
   known_failures = known_szs_status_failures,
47981
df35a8dd6368 gracefully handle definition-looking premises
blanchet
parents: 47976
diff changeset
   397
   prem_role = Hypothesis,
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   398
   best_slices =
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   399
     (* FUDGE *)
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   400
     K [(1.0, (((150, ""), THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), ""))],
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   401
   best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *),
53515
f5b678b155f6 adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents: 53225
diff changeset
   402
   best_max_new_mono_instances = default_max_new_mono_instances}
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   403
47646
9460f3f22365 tried to make SML/NJ happy
blanchet
parents: 47606
diff changeset
   404
val satallax = (satallaxN, fn () => satallax_config)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   405
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   406
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   407
(* SPASS *)
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42723
diff changeset
   408
48005
eeede26f2721 killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents: 48004
diff changeset
   409
val spass_H1SOS = "-Heuristic=1 -SOS"
50333
20c69b00e73c tweak SPASS default a tiny bit, so that a more interesting heuristic is chosen when "slicing=false" (for experiments)
blanchet
parents: 49991
diff changeset
   410
val spass_H2 = "-Heuristic=2"
48005
eeede26f2721 killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents: 48004
diff changeset
   411
val spass_H2LR0LT0 = "-Heuristic=2 -LR=0 -LT=0"
eeede26f2721 killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents: 48004
diff changeset
   412
val spass_H2NuVS0 = "-Heuristic=2 -RNuV=1 -Sorts=0"
eeede26f2721 killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents: 48004
diff changeset
   413
val spass_H2NuVS0Red2 = "-Heuristic=2 -RNuV=1 -Sorts=0 -RFRew=2 -RBRew=2 -RTaut=2"
50333
20c69b00e73c tweak SPASS default a tiny bit, so that a more interesting heuristic is chosen when "slicing=false" (for experiments)
blanchet
parents: 49991
diff changeset
   414
val spass_H2SOS = "-Heuristic=2 -SOS"
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   415
48005
eeede26f2721 killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents: 48004
diff changeset
   416
val spass_config : atp_config =
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   417
  let
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   418
    val format = DFG Monomorphic
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   419
  in
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   420
    {exec = (["SPASS_HOME"], ["SPASS"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   421
     arguments = fn _ => fn full_proofs => fn extra_options => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   422
       ["-Isabelle=1 " ^ (if full_proofs then "-CNFRenaming=0 -Splits=0 " else "") ^
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   423
        "-TimeLimit=" ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   424
        |> extra_options <> "" ? prefix (extra_options ^ " ")],
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   425
     proof_delims = [("Here is a proof", "Formulae used in the proof")],
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   426
     known_failures =
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   427
       [(GaveUp, "SPASS beiseite: Completion found"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   428
        (TimedOut, "SPASS beiseite: Ran out of time"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   429
        (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   430
        (MalformedInput, "Undefined symbol"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   431
        (MalformedInput, "Free Variable"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   432
        (Unprovable, "No formulae and clauses found in input file"),
73436
e92f2e44e4d8 removed spurious references to perl / libwww-perl;
wenzelm
parents: 73435
diff changeset
   433
        (InternalError, "Please report this error")],
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   434
     prem_role = Conjecture,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   435
     best_slices = fn _ =>
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   436
       (* FUDGE *)
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   437
       [(0.1667, (((150, meshN), format, "mono_native", combsN, true), "")),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   438
        (0.1667, (((500, meshN), format, "mono_native", liftingN, true), spass_H2SOS)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   439
        (0.1666, (((50, meshN), format,  "mono_native", liftingN, true), spass_H2LR0LT0)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   440
        (0.1000, (((250, meshN), format, "mono_native", combsN, true), spass_H2NuVS0)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   441
        (0.1000, (((1000, mepoN), format, "mono_native", liftingN, true), spass_H1SOS)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   442
        (0.1000, (((150, meshN), format, "poly_guards??", liftingN, false), spass_H2NuVS0Red2)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   443
        (0.1000, (((300, meshN), format, "mono_native", combsN, true), spass_H2SOS)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   444
        (0.1000, (((100, meshN), format, "mono_native", combs_and_liftingN, true), spass_H2))],
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   445
     best_max_mono_iters = default_max_mono_iters,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   446
     best_max_new_mono_instances = default_max_new_mono_instances}
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   447
  end
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   448
48005
eeede26f2721 killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents: 48004
diff changeset
   449
val spass = (spassN, fn () => spass_config)
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   450
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   451
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   452
(* Vampire *)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   453
74311
19022ea3f8cc clarified name and options for old vampire-4.2.2;
wenzelm
parents: 74207
diff changeset
   454
val vampire_basic_options =
19022ea3f8cc clarified name and options for old vampire-4.2.2;
wenzelm
parents: 74207
diff changeset
   455
  "--proof tptp --output_axiom_names on" ^
19022ea3f8cc clarified name and options for old vampire-4.2.2;
wenzelm
parents: 74207
diff changeset
   456
  (if ML_System.platform_is_windows
19022ea3f8cc clarified name and options for old vampire-4.2.2;
wenzelm
parents: 74207
diff changeset
   457
   then ""  (*time slicing is not support in the Windows version of Vampire*)
19022ea3f8cc clarified name and options for old vampire-4.2.2;
wenzelm
parents: 74207
diff changeset
   458
   else " --mode casc")
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   459
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   460
val vampire_full_proof_options =
71793
e771b8157fc7 tweaked Vampire's options + tuning
blanchet
parents: 71353
diff changeset
   461
  " --proof_extra free --forced_options avatar=off:equality_proxy=off:general_splitting=off:inequality_splitting=0:naming=0"
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   462
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   463
val vampire_config : atp_config =
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   464
  let
74351
d8dbe7525ff1 enabled FOOL for Vampire in Sledgehammer
desharna
parents: 74350
diff changeset
   465
    val format = TFF (With_FOOL, Monomorphic)
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   466
  in
74350
398b7bb9ebdd used Vampire 4.5.1 in Sledgehammer
desharna
parents: 74311
diff changeset
   467
    {exec = (["VAMPIRE_HOME"], ["vampire"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   468
     arguments = fn _ => fn full_proofs => fn sos => fn timeout => fn problem => fn _ =>
74352
fb8ce6090437 removed checks for non-commercial usage of Vampire as it is now under BSD licence
desharna
parents: 74351
diff changeset
   469
       [vampire_basic_options ^ (if full_proofs then " " ^ vampire_full_proof_options else "") ^
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   470
         " -t " ^ string_of_int (to_secs 1 timeout) ^ " --input_file " ^ File.bash_path problem
74352
fb8ce6090437 removed checks for non-commercial usage of Vampire as it is now under BSD licence
desharna
parents: 74351
diff changeset
   471
         |> sos = sosN ? prefix "--sos on "],
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   472
     proof_delims =
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   473
       [("=========== Refutation ==========",
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   474
         "======= End of refutation =======")] @
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   475
       tstp_proof_delims,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   476
     known_failures =
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   477
       [(GaveUp, "UNPROVABLE"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   478
        (GaveUp, "CANNOT PROVE"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   479
        (Unprovable, "Satisfiability detected"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   480
        (Unprovable, "Termination reason: Satisfiable"),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   481
        (Interrupted, "Aborted by signal SIGINT")] @
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   482
       known_szs_status_failures,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   483
     prem_role = Hypothesis,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   484
     best_slices = fn ctxt =>
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   485
       (* FUDGE *)
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   486
       [(0.333, (((500, meshN), format, "mono_native", combs_or_liftingN, false), sosN)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   487
        (0.333, (((150, meshN), format, "poly_tags??", combs_or_liftingN, false), sosN)),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   488
        (0.334, (((50, meshN), format, "mono_native", combs_or_liftingN, false), no_sosN))]
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   489
       |> Config.get ctxt force_sos ? (hd #> apfst (K 1.0) #> single),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   490
     best_max_mono_iters = default_max_mono_iters,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   491
     best_max_new_mono_instances = 2 * default_max_new_mono_instances (* FUDGE *)}
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   492
  end
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   493
47646
9460f3f22365 tried to make SML/NJ happy
blanchet
parents: 47606
diff changeset
   494
val vampire = (vampireN, fn () => vampire_config)
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   495
68563
05fb05f94686 added option for noncommercial Vampire
blanchet
parents: 68328
diff changeset
   496
48803
ffa31bf5c662 tone down "z3_tptp", now that Z3 (starting with 4.1) no longer supports TPTP TFF0
blanchet
parents: 48801
diff changeset
   497
(* Z3 with TPTP syntax (half experimental, half legacy) *)
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   498
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
   499
val z3_tptp_config : atp_config =
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   500
  let
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   501
    val format = TFF (Without_FOOL, Monomorphic)
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   502
  in
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   503
    {exec = (["Z3_TPTP_HOME"], ["z3_tptp"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   504
     arguments = fn _ => fn _ => fn _ => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   505
       ["-proof -t:" ^ string_of_int (to_secs 1 timeout) ^ " -file:" ^ File.bash_path problem],
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   506
     proof_delims = [("SZS status Theorem", "")],
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   507
     known_failures = known_szs_status_failures,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   508
     prem_role = Hypothesis,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   509
     best_slices =
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   510
       (* FUDGE *)
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   511
       K [(0.5, (((250, meshN), format, "mono_native", combsN, false), "")),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   512
          (0.25, (((125, mepoN), format, "mono_native", combsN, false), "")),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   513
          (0.125, (((62, mashN), format, "mono_native", combsN, false), "")),
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   514
          (0.125, (((31, meshN), format, "mono_native", combsN, false), ""))],
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   515
     best_max_mono_iters = default_max_mono_iters,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   516
     best_max_new_mono_instances = 2 * default_max_new_mono_instances (* FUDGE *)}
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   517
  end
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   518
47646
9460f3f22365 tried to make SML/NJ happy
blanchet
parents: 47606
diff changeset
   519
val z3_tptp = (z3_tptpN, fn () => z3_tptp_config)
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   520
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   521
69717
eb74ff534b27 tune ATP settings
blanchet
parents: 69593
diff changeset
   522
(* Zipperposition *)
eb74ff534b27 tune ATP settings
blanchet
parents: 69593
diff changeset
   523
72174
585b877df698 basic integration of Zipperposition 2.0
blanchet
parents: 71793
diff changeset
   524
val zipperposition_blsimp = "--mode ho-pragmatic --max-inferences 3 --ho-max-app-projections 0 --ho-max-elims 0 --ho-max-rigid-imitations 2 --ho-max-identifications 0 --ho-unif-max-depth 2 --boolean-reasoning no-cases --ext-rules ext-family --ext-rules-max-depth 1 --kbo-weight-fun invdocc --ho-prim-enum tf --ho-prim-enum-early-bird true --tptp-def-as-rewrite --ho-unif-level pragmatic-framework -q '1|const|conjecture-relative-var(1,s,f)' -q '1|prefer-processed|pnrefined(1,1,1,2,2,2,0.5)' -q '1|prefer-sos|staggered(1)' -q '2|prefer-fo|default' -q '1|prefer-neg-unit|orient-lmax(2,1,2,1,1)' -q '2|prefer-easy-ho|conjecture-relative-struct(1.5,3.5,2,3)' --ho-elim-leibniz 2 --ho-fixpoint-decider true --ho-pattern-decider false --ho-solid-decider true --ho-max-solidification 12  --select e-selection11 --solve-formulas true --sup-at-vars false --sup-at-var-headed false --lazy-cnf true --lazy-cnf-kind simp --lazy-cnf-renaming-threshold 4 --sine 50 --sine-tolerance 1.7 --sine-depth-max 3 --sine-depth-min 1 --sine-trim-implications true --ho-selection-restriction none --sup-from-var-headed false --sine-trim-implications true"
585b877df698 basic integration of Zipperposition 2.0
blanchet
parents: 71793
diff changeset
   525
val zipperposition_s6 = "--tptp-def-as-rewrite --rewrite-before-cnf true --mode ho-competitive --boolean-reasoning no-cases --ext-rules off --ho-prim-enum none --recognize-injectivity true --ho-elim-leibniz off --ho-unif-level full-framework --no-max-vars -q '3|const|conjecture-relative-var(1.02,l,f)' -q '1|prefer-ho-steps|conjecture-relative-var(1,s,f)' -q '1|prefer-processed|fifo' -q '3|by-app-var-num|pnrefined(2,1,1,1,2,2,2)' --select ho-selection5 --prec-gen-fun unary_first --solid-subsumption false --ignore-orphans false --ho-solid-decider true --ho-fixpoint-decider true --ho-pattern-decider true --sup-at-vars false --sup-at-var-headed false --sup-from-var-headed false --ho-neg-ext-simpl true"
585b877df698 basic integration of Zipperposition 2.0
blanchet
parents: 71793
diff changeset
   526
val zipperposition_cdots = "--mode ho-competitive --boolean-reasoning cases-simpl --ext-rules ext-family --ext-rules-max-depth 1 --ho-prim-enum pragmatic --ho-prim-max 1 --bool-subterm-selection A --avatar off --recognize-injectivity true --ho-elim-leibniz 1 --ho-unif-level full-framework --no-max-vars -q '6|prefer-sos|pnrefined(1,1,1,2,2,2,0.5)' -q '6|const|conjecture-relative-var(1.02,l,f)' -q '1|prefer-processed|fifo' -q '1|prefer-non-goals|conjecture-relative-var(1,l,f)' -q '4|prefer-easy-ho|conjecture-relative-var(1.01,s,f)' --select e-selection7 --ho-choice-inst true --sine 50 --sine-tolerance 2 --sine-depth-max 4 --sine-depth-min 1 --scan-clause-ac true --lambdasup 0 --kbo-weight-fun invfreqrank"
585b877df698 basic integration of Zipperposition 2.0
blanchet
parents: 71793
diff changeset
   527
57154
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57008
diff changeset
   528
val zipperposition_config : atp_config =
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   529
  let
74206
9c6159cbf9ee disabled 'ite' in Zipperposition until we upgrade to a version of Zip that supports it and we generate the proper syntax
blanchet
parents: 74117
diff changeset
   530
    val format = THF (Without_FOOL, Polymorphic, THF_Without_Choice)
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   531
  in
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   532
    {exec = (["ZIPPERPOSITION_HOME"], ["zipperposition"]),
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   533
     arguments = fn _ => fn _ => fn extra_options => fn timeout => fn problem => fn _ =>
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   534
       ["--input tptp --output tptp --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^ File.bash_path problem
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   535
        |> extra_options <> "" ? prefix (extra_options ^ " ")],
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   536
     proof_delims = tstp_proof_delims,
74207
adf767b94f77 handle Zipperposition's ResourceOut gracefully
blanchet
parents: 74206
diff changeset
   537
     known_failures =
adf767b94f77 handle Zipperposition's ResourceOut gracefully
blanchet
parents: 74206
diff changeset
   538
       [(TimedOut, "SZS status ResourceOut")] @   (* odd way of timing out *)
adf767b94f77 handle Zipperposition's ResourceOut gracefully
blanchet
parents: 74206
diff changeset
   539
       known_szs_status_failures,
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   540
     prem_role = Hypothesis,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   541
     best_slices = fn _ =>
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   542
       (* FUDGE *)
74206
9c6159cbf9ee disabled 'ite' in Zipperposition until we upgrade to a version of Zip that supports it and we generate the proper syntax
blanchet
parents: 74117
diff changeset
   543
       [(0.333, (((128, "meshN"), format, "mono_native_higher", keep_lamsN, false), zipperposition_blsimp)),
9c6159cbf9ee disabled 'ite' in Zipperposition until we upgrade to a version of Zip that supports it and we generate the proper syntax
blanchet
parents: 74117
diff changeset
   544
        (0.333, (((32, "meshN"), format, "poly_native_higher", keep_lamsN, false), zipperposition_s6)),
9c6159cbf9ee disabled 'ite' in Zipperposition until we upgrade to a version of Zip that supports it and we generate the proper syntax
blanchet
parents: 74117
diff changeset
   545
        (0.334, (((512, "meshN"), format, "mono_native_higher", keep_lamsN, false), zipperposition_cdots))],
73375
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   546
     best_max_mono_iters = default_max_mono_iters,
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   547
     best_max_new_mono_instances = default_max_new_mono_instances}
a80fd78c85bd tuned best_slices in atp_config
desharna
parents: 73374
diff changeset
   548
  end
57154
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57008
diff changeset
   549
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57008
diff changeset
   550
val zipperposition = (zipperpositionN, fn () => zipperposition_config)
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57008
diff changeset
   551
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57008
diff changeset
   552
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   553
(* Remote ATP invocation via SystemOnTPTP *)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   554
73426
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   555
val no_remote_systems = {url = "", systems = [] : string list}
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   556
val remote_systems = Synchronized.var "atp_remote_systems" no_remote_systems
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   557
49984
blanchet
parents: 48803
diff changeset
   558
fun get_remote_systems () =
73426
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   559
  Timeout.apply (seconds 10.0) SystemOnTPTP.list_systems ()
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   560
    handle ERROR msg => (warning msg; no_remote_systems)
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   561
      | Timeout.TIMEOUT _ => no_remote_systems
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   562
49984
blanchet
parents: 48803
diff changeset
   563
fun find_remote_system name [] systems =
42537
25ceb855a18b improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
blanchet
parents: 42535
diff changeset
   564
    find_first (String.isPrefix (name ^ "---")) systems
49984
blanchet
parents: 48803
diff changeset
   565
  | find_remote_system name (version :: versions) systems =
38690
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   566
    case find_first (String.isPrefix (name ^ "---" ^ version)) systems of
49984
blanchet
parents: 48803
diff changeset
   567
      NONE => find_remote_system name versions systems
38690
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   568
    | res => res
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   569
49984
blanchet
parents: 48803
diff changeset
   570
fun get_remote_system name versions =
73426
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   571
  Synchronized.change_result remote_systems (fn remote =>
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   572
    (if #url remote <> SystemOnTPTP.get_url () orelse null (#systems remote)
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   573
      then get_remote_systems () else remote) |> ` #systems)
bd8bce50b9d4 use SystemOnTPTP.list_systems from Isabelle/Scala, with dynamic URL option and more elementary error messages;
wenzelm
parents: 73375
diff changeset
   574
  |> `(find_remote_system name versions)
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
   575
49984
blanchet
parents: 48803
diff changeset
   576
fun the_remote_system name versions =
54788
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   577
  (case get_remote_system name versions of
42955
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   578
    (SOME sys, _) => sys
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63116
diff changeset
   579
  | (NONE, []) => error "SystemOnTPTP is currently not available"
42955
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   580
  | (NONE, syss) =>
54788
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   581
    (case syss |> filter_out (String.isPrefix "%") |> filter_out (curry (op =) "") of
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63116
diff changeset
   582
      [] => error "SystemOnTPTP is currently not available"
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63116
diff changeset
   583
    | [msg] => error ("SystemOnTPTP is currently not available: " ^ msg)
46480
24990fae5f92 better error message
blanchet
parents: 46455
diff changeset
   584
    | syss =>
54788
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   585
      error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n(Available systems: " ^
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   586
        commas_quote syss ^ ".)")))
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   587
72174
585b877df698 basic integration of Zipperposition 2.0
blanchet
parents: 71793
diff changeset
   588
val max_remote_secs = 1000   (* give Geoff Sutcliffe's servers a break *)
41148
f5229ab54284 added timeout max for remote server invocation
blanchet
parents: 40669
diff changeset
   589
73435
1cc848548f21 invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents: 73432
diff changeset
   590
val isabelle_scala_function = (["SCALA_HOME"], ["bin/scala"])
1cc848548f21 invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents: 73432
diff changeset
   591
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   592
fun remote_config system_name system_versions proof_delims known_failures prem_role best_slice =
73435
1cc848548f21 invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents: 73432
diff changeset
   593
  {exec = isabelle_scala_function,
73432
3dcca6c4e5cc clarified signature: more explicit types;
wenzelm
parents: 73426
diff changeset
   594
   arguments = fn _ => fn _ => fn command => fn timeout => fn problem => fn _ =>
73435
1cc848548f21 invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents: 73432
diff changeset
   595
     [the_remote_system system_name system_versions,
1cc848548f21 invoke remote ATP via SystemOnTPTP.run_systems from Isabelle/Scala (without perl);
wenzelm
parents: 73432
diff changeset
   596
      Isabelle_System.absolute_path problem,
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   597
      command, string_of_int (Int.min (max_remote_secs, to_secs 1 timeout) * 1000)],
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42955
diff changeset
   598
   proof_delims = union (op =) tstp_proof_delims proof_delims,
73436
e92f2e44e4d8 removed spurious references to perl / libwww-perl;
wenzelm
parents: 73435
diff changeset
   599
   known_failures = known_failures @ known_says_failures,
47976
6b13451135a9 tuned names
blanchet
parents: 47974
diff changeset
   600
   prem_role = prem_role,
48716
1d2a12bb0640 stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents: 48715
diff changeset
   601
   best_slices = fn ctxt => [(1.0, best_slice ctxt)],
47962
137883567114 lower the monomorphization thresholds for less scalable provers
blanchet
parents: 47955
diff changeset
   602
   best_max_mono_iters = default_max_mono_iters,
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   603
   best_max_new_mono_instances = default_max_new_mono_instances} : atp_config
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42332
diff changeset
   604
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   605
fun remotify_config system_name system_versions best_slice
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   606
    ({proof_delims, known_failures, prem_role, ...} : atp_config) =
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   607
  remote_config system_name system_versions proof_delims known_failures prem_role best_slice
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   608
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   609
fun remote_atp name system_name system_versions proof_delims known_failures prem_role best_slice =
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   610
  (remote_prefix ^ name, fn () =>
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   611
     remote_config system_name system_versions proof_delims known_failures prem_role best_slice)
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   612
fun remotify_atp (name, config) system_name system_versions best_slice =
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   613
  (remote_prefix ^ name, remotify_config system_name system_versions best_slice o config)
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
   614
57269
1df6f747f164 changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
blanchet
parents: 57265
diff changeset
   615
fun gen_remote_waldmeister name type_enc =
57265
cab38f7a3adb use right delimiters for Waldmeister proofs
blanchet
parents: 57264
diff changeset
   616
  remote_atp name "Waldmeister" ["710"] tstp_proof_delims
cab38f7a3adb use right delimiters for Waldmeister proofs
blanchet
parents: 57264
diff changeset
   617
    ([(OutOfResources, "Too many function symbols"),
cab38f7a3adb use right delimiters for Waldmeister proofs
blanchet
parents: 57264
diff changeset
   618
      (Inappropriate, "****  Unexpected end of file."),
cab38f7a3adb use right delimiters for Waldmeister proofs
blanchet
parents: 57264
diff changeset
   619
      (Crashed, "Unrecoverable Segmentation Fault")]
cab38f7a3adb use right delimiters for Waldmeister proofs
blanchet
parents: 57264
diff changeset
   620
     @ known_szs_status_failures)
57264
13db1d078743 added 'waldmeister_new' as ATP
blanchet
parents: 57262
diff changeset
   621
    Hypothesis
57269
1df6f747f164 changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
blanchet
parents: 57265
diff changeset
   622
    (K (((50, ""), CNF_UEQ, type_enc, combsN, false), "") (* FUDGE *))
57264
13db1d078743 added 'waldmeister_new' as ATP
blanchet
parents: 57262
diff changeset
   623
52094
018cc7c7e640 updated remote provers
blanchet
parents: 52073
diff changeset
   624
val remote_agsyhol =
018cc7c7e640 updated remote provers
blanchet
parents: 52073
diff changeset
   625
  remotify_atp agsyhol "agsyHOL" ["1.0", "1"]
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   626
    (K (((60, ""), THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), "") (* FUDGE *))
70937
fe114eca312e added support for repote Alt-Ergo
blanchet
parents: 70935
diff changeset
   627
val remote_alt_ergo =
fe114eca312e added support for repote Alt-Ergo
blanchet
parents: 70935
diff changeset
   628
  remotify_atp alt_ergo "Alt-Ergo" ["0.95.2"]
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   629
    (K (((250, ""), TFF (Without_FOOL, Polymorphic), "poly_native", keep_lamsN, false), "") (* FUDGE *))
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   630
val remote_e =
63768
a09cfea0c2c9 adapted remote E
blanchet
parents: 63692
diff changeset
   631
  remotify_atp e "E" ["2.0", "1.9.1", "1.8"]
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   632
    (K (((750, ""), TFF (Without_FOOL, Monomorphic), "mono_native", combsN, false), "") (* FUDGE *))
48700
d06138bfeb45 added iProver(-Eq) local
blanchet
parents: 48653
diff changeset
   633
val remote_iprover =
52094
018cc7c7e640 updated remote provers
blanchet
parents: 52073
diff changeset
   634
  remotify_atp iprover "iProver" ["0.99"]
58084
9f77084444df pass options to remote Vampire
blanchet
parents: 57759
diff changeset
   635
    (K (((150, ""), FOF, "mono_guards??", liftingN, false), "") (* FUDGE *))
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   636
val remote_leo2 =
52094
018cc7c7e640 updated remote provers
blanchet
parents: 52073
diff changeset
   637
  remotify_atp leo2 "LEO-II" ["1.5.0", "1.4", "1.3", "1.2", "1"]
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   638
    (K (((40, ""), THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", liftingN, false), "") (* FUDGE *))
67021
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   639
val remote_leo3 =
41f1f8c4259b integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents: 66544
diff changeset
   640
  remotify_atp leo3 "Leo-III" ["1.1"]
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   641
    (K (((150, ""), THF (Without_FOOL, Polymorphic, THF_Without_Choice), "poly_native_higher", keep_lamsN, false), "") (* FUDGE *))
57269
1df6f747f164 changed type encoding for new Waldmeister, to trigger filtering of 'dangerous' lemmas
blanchet
parents: 57265
diff changeset
   642
val remote_waldmeister = gen_remote_waldmeister waldmeisterN "raw_mono_tags??"
70940
ce3a05ad07b7 added support for Zipperposition on SystemOnTPTP
blanchet
parents: 70939
diff changeset
   643
val remote_zipperposition =
72174
585b877df698 basic integration of Zipperposition 2.0
blanchet
parents: 71793
diff changeset
   644
  remotify_atp zipperposition "Zipperpin" ["2.0"]
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   645
    (K (((512, ""), THF (Without_FOOL, Monomorphic, THF_Without_Choice), "mono_native_higher", keep_lamsN, false), "") (* FUDGE *))
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   646
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   647
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   648
(* Dummy prover *)
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   649
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   650
fun dummy_config prem_role format type_enc uncurried_aliases : atp_config =
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   651
  {exec = (["ISABELLE_ATP"], ["scripts/dummy_atp"]),
73568
bdba138d462d clarified signature: more structured arguments, notably for remote provers;
wenzelm
parents: 73437
diff changeset
   652
   arguments = K (K (K (K (K (K []))))),
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   653
   proof_delims = [],
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   654
   known_failures = known_szs_status_failures,
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   655
   prem_role = prem_role,
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   656
   best_slices =
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   657
     K [(1.0, (((200, ""), format, type_enc,
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   658
                if is_format_higher_order format then keep_lamsN
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   659
                else combsN, uncurried_aliases), ""))],
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   660
   best_max_mono_iters = default_max_mono_iters,
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   661
   best_max_new_mono_instances = default_max_new_mono_instances}
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   662
74117
30ab39ab4117 added dummy_fof prover to Sledgehammer
desharna
parents: 74109
diff changeset
   663
30ab39ab4117 added dummy_fof prover to Sledgehammer
desharna
parents: 74109
diff changeset
   664
val dummy_fof_format = FOF
30ab39ab4117 added dummy_fof prover to Sledgehammer
desharna
parents: 74109
diff changeset
   665
val dummy_fof_config = dummy_config Hypothesis dummy_fof_format "mono_guards??" false
30ab39ab4117 added dummy_fof prover to Sledgehammer
desharna
parents: 74109
diff changeset
   666
val dummy_fof = (dummy_fofN, fn () => dummy_fof_config)
30ab39ab4117 added dummy_fof prover to Sledgehammer
desharna
parents: 74109
diff changeset
   667
72588
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   668
val dummy_tfx_format = TFF (With_FOOL, Polymorphic)
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   669
val dummy_tfx_config = dummy_config Hypothesis dummy_tfx_format "mono_native_fool" false
c7e2a9bdc585 Added support for TFX to Sledgehammer
desharna
parents: 72403
diff changeset
   670
val dummy_tfx = (dummy_tfxN, fn () => dummy_tfx_config)
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   671
74109
ed1f576df9c4 added dummy_thf prover to Sledgehammer
desharna
parents: 74047
diff changeset
   672
val dummy_thf_format = THF (With_FOOL, Polymorphic, THF_With_Choice)
ed1f576df9c4 added dummy_thf prover to Sledgehammer
desharna
parents: 74047
diff changeset
   673
val dummy_thf_config = dummy_config Hypothesis dummy_thf_format "mono_native_higher_fool" false
ed1f576df9c4 added dummy_thf prover to Sledgehammer
desharna
parents: 74047
diff changeset
   674
val dummy_thf = (dummy_thfN, fn () => dummy_thf_config)
ed1f576df9c4 added dummy_thf prover to Sledgehammer
desharna
parents: 74047
diff changeset
   675
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   676
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   677
(* Setup *)
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   678
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   679
fun add_atp (name, config) thy =
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   680
  Data.map (Symtab.update_new (name, (config, stamp ()))) thy
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63116
diff changeset
   681
  handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name)
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   682
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   683
fun get_atp thy name =
54788
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   684
  fst (the (Symtab.lookup (Data.get thy) name))
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 63116
diff changeset
   685
  handle Option.Option => error ("Unknown ATP: " ^ name)
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   686
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41725
diff changeset
   687
val supported_atps = Symtab.keys o Data.get
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   688
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   689
fun is_atp_installed thy name =
48376
416e4123baf3 use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents: 48232
diff changeset
   690
  let val {exec, ...} = get_atp thy name () in
73374
316e12147698 tuned exec field in atp_config
desharna
parents: 73294
diff changeset
   691
    exists (fn var => getenv var <> "") (fst exec)
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   692
  end
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   693
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   694
fun refresh_systems_on_tptp () =
49984
blanchet
parents: 48803
diff changeset
   695
  Synchronized.change remote_systems (fn _ => get_remote_systems ())
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   696
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   697
fun effective_term_order ctxt atp =
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   698
  let val ord = Config.get ctxt term_order in
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   699
    if ord = smartN then
54788
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   700
      {is_lpo = false, gen_weights = (atp = spassN), gen_prec = (atp = spassN),
72401
2783779b7dd3 removed obsolete unmaintained experimental prover Pirate
blanchet
parents: 72400
diff changeset
   701
       gen_simp = false}
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   702
    else
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   703
      let val is_lpo = String.isSubstring lpoN ord in
54788
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   704
        {is_lpo = is_lpo, gen_weights = not is_lpo andalso String.isSubstring xweightsN ord,
a898e15b522a primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents: 54197
diff changeset
   705
         gen_prec = String.isSubstring xprecN ord, gen_simp = String.isSubstring xsimpN ord}
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   706
      end
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   707
  end
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   708
52073
ccb292952774 started adding agsyHOL as an experimental prover
blanchet
parents: 51998
diff changeset
   709
val atps =
72403
4a3169d8885c removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents: 72401
diff changeset
   710
  [agsyhol, alt_ergo, e, iprover, leo2, leo3, satallax, spass, vampire, z3_tptp, zipperposition,
4a3169d8885c removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents: 72401
diff changeset
   711
   remote_agsyhol, remote_alt_ergo, remote_e, remote_iprover, remote_leo2, remote_leo3,
74117
30ab39ab4117 added dummy_fof prover to Sledgehammer
desharna
parents: 74109
diff changeset
   712
   remote_waldmeister, remote_zipperposition, dummy_fof, dummy_tfx, dummy_thf]
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47053
diff changeset
   713
57262
b2c629647a14 moved code around
blanchet
parents: 57261
diff changeset
   714
val _ = Theory.setup (fold add_atp atps)
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
   715
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
   716
end;