src/HOL/Tools/ATP/atp_systems.ML
author blanchet
Fri, 03 Feb 2012 18:00:55 +0100
changeset 46407 30e9720cc0b9
parent 46402 ef8d65f64f77
child 46409 d4754183ccce
permissions -rw-r--r--
optimization: slice caching in case two consecutive slices are nearly identical
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
36376
e83d52a52449 renamed module "ATP_Wrapper" to "ATP_Systems"
blanchet
parents: 36371
diff changeset
     8
signature ATP_SYSTEMS =
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
     9
sig
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
    10
  type atp_format = ATP_Problem.atp_format
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
    11
  type formula_kind = ATP_Problem.formula_kind
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
    12
  type failure = ATP_Proof.failure
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    13
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    14
  type atp_config =
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    15
    {exec : string * string,
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    16
     required_execs : (string * string) list,
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    17
     arguments :
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
    18
       Proof.context -> bool -> string -> Time.time
43354
396aaa15dd8b pass --trim option to "eproof" script to speed up proof reconstruction
blanchet
parents: 43288
diff changeset
    19
       -> (unit -> (string * real) list) -> string,
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,
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    21
     known_failures : (failure * string) list,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
    22
     conj_sym_kind : formula_kind,
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
    23
     prem_kind : formula_kind,
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
    24
     best_slices :
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
    25
       Proof.context
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
    26
       -> (real * (bool * ((int * atp_format * string * string) * string)))
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
    27
            list}
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    28
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
    29
  val force_sos : bool 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
    30
  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
    31
  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
    32
  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
    33
  val e_sym_offset_weightN : string
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    34
  val e_weight_method : string Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    35
  val e_default_fun_weight : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    36
  val e_fun_weight_base : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    37
  val e_fun_weight_span : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    38
  val e_default_sym_offs_weight : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    39
  val e_sym_offs_weight_base : real Config.T
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    40
  val e_sym_offs_weight_span : real Config.T
46380
7e049e9f5c8b new SPASS setup
blanchet
parents: 46370
diff changeset
    41
  val spass_incompleteN : string
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    42
  val eN : string
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
    43
  val e_sineN : string
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
    44
  val e_tofofN : string
45338
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
    45
  val iproverN : string
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
    46
  val iprover_eqN : string
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
    47
  val leo2N : string
45365
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
    48
  val dummy_tff1N : string
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
    49
  val dummy_thfN : string
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
    50
  val satallaxN : string
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
    51
  val snarkN : string
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    52
  val spassN : string
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
    53
  val spass_newN : string
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    54
  val vampireN : string
42938
c124032ac96f added Waldmeister
blanchet
parents: 42937
diff changeset
    55
  val waldmeisterN : string
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
    56
  val z3_tptpN : string
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
    57
  val remote_prefix : string
41738
eb98c60a6cf0 added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
blanchet
parents: 41727
diff changeset
    58
  val remote_atp :
eb98c60a6cf0 added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
blanchet
parents: 41727
diff changeset
    59
    string -> string -> string list -> (string * string) list
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
    60
    -> (failure * string) list -> formula_kind -> formula_kind
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
    61
    -> (Proof.context -> int * atp_format * string * string)
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
    62
    -> string * atp_config
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    63
  val add_atp : string * atp_config -> theory -> theory
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    64
  val get_atp : theory -> string -> atp_config
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41725
diff changeset
    65
  val supported_atps : theory -> string list
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    66
  val is_atp_installed : theory -> string -> bool
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
    67
  val refresh_systems_on_tptp : unit -> unit
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
    68
  val setup : theory -> theory
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    69
end;
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    70
36376
e83d52a52449 renamed module "ATP_Wrapper" to "ATP_Systems"
blanchet
parents: 36371
diff changeset
    71
structure ATP_Systems : ATP_SYSTEMS =
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    72
struct
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
    73
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
    74
open ATP_Problem
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
    75
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45876
diff changeset
    76
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
    77
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    78
(* ATP configuration *)
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
    79
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
    80
type atp_config =
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    81
  {exec : string * string,
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    82
   required_execs : (string * string) list,
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
    83
   arguments :
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
    84
     Proof.context -> bool -> string -> Time.time
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
    85
     -> (unit -> (string * real) list) -> string,
42578
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    86
   proof_delims : (string * string) list,
1eaf4d437d4c define type system in ATP module so that ATPs can suggest a type system
blanchet
parents: 42577
diff changeset
    87
   known_failures : (failure * string) list,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
    88
   conj_sym_kind : formula_kind,
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
    89
   prem_kind : formula_kind,
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
    90
   best_slices :
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
    91
     Proof.context
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
    92
     -> (real * (bool * ((int * atp_format * string * string) * string))) list}
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
    93
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
    94
(* "best_slices" must be found empirically, taking a wholistic approach since
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
    95
   the ATPs are run in parallel. The "real" component gives the faction of 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
    96
   time available given to the slice and should add up to 1.0. The "bool"
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
    97
   component indicates whether the slice's strategy is complete; the "int", 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
    98
   preferred number of facts to pass; the first "string", the preferred type
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
    99
   system (which should be sound or quasi-sound); the second "string", the
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   100
   preferred lambda translation scheme; the third "string", extra information to
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   101
   the prover (e.g., SOS or no SOS).
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
   102
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
   103
   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
   104
   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
   105
   slicing is disabled (e.g., by the minimizer). *)
42710
84fcce345b5d further improved type system setup based on Judgment Days
blanchet
parents: 42709
diff changeset
   106
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   107
val known_perl_failures =
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38092
diff changeset
   108
  [(CantConnect, "HTTP error"),
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38092
diff changeset
   109
   (NoPerl, "env: perl"),
38065
9069e1ad1527 improved ATP error handling some more
blanchet
parents: 38064
diff changeset
   110
   (NoLibwwwPerl, "Can't locate HTTP")]
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
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
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   130
(* named ATPs *)
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   131
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   132
val eN = "e"
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   133
val e_sineN = "e_sine"
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   134
val e_tofofN = "e_tofof"
45338
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
   135
val iproverN = "iprover"
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
   136
val iprover_eqN = "iprover_eq"
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   137
val leo2N = "leo2"
45365
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   138
val dummy_tff1N = "dummy_tff1" (* experimental *)
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   139
val dummy_thfN = "dummy_thf" (* experimental *)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   140
val satallaxN = "satallax"
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   141
val snarkN = "snark"
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   142
val spassN = "spass"
45365
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   143
val spass_newN = "spass_new" (* experimental *)
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   144
val vampireN = "vampire"
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   145
val waldmeisterN = "waldmeister"
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
   146
val z3_tptpN = "z3_tptp"
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   147
val remote_prefix = "remote_"
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   148
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   149
structure Data = Theory_Data
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   150
(
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   151
  type T = (atp_config * stamp) Symtab.table
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   152
  val empty = Symtab.empty
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   153
  val extend = I
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   154
  fun merge data : T =
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   155
    Symtab.merge (eq_snd (op =)) data
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   156
    handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   157
)
38017
3ad3e3ca2451 move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents: 38015
diff changeset
   158
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
   159
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
   160
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   161
val sosN = "sos"
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   162
val no_sosN = "no_sos"
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   163
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   164
val force_sos = Attrib.setup_config_bool @{binding atp_force_sos} (K false)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   165
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   166
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   167
(* E *)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   168
44420
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   169
fun is_old_e_version () = (string_ord (getenv "E_VERSION", "1.2w") = LESS)
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   170
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   171
val tstp_proof_delims =
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42955
diff changeset
   172
  [("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation"),
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42955
diff changeset
   173
   ("% SZS output start CNFRefutation", "% SZS output end CNFRefutation")]
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   174
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   175
val e_smartN = "smart"
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   176
val e_autoN = "auto"
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   177
val e_fun_weightN = "fun_weight"
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   178
val e_sym_offset_weightN = "sym_offset_weight"
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   179
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   180
val e_weight_method =
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   181
  Attrib.setup_config_string @{binding atp_e_weight_method} (K e_smartN)
41770
a710e96583d5 adjust fudge factors
blanchet
parents: 41769
diff changeset
   182
(* FUDGE *)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   183
val e_default_fun_weight =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   184
  Attrib.setup_config_real @{binding atp_e_default_fun_weight} (K 20.0)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   185
val e_fun_weight_base =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   186
  Attrib.setup_config_real @{binding atp_e_fun_weight_base} (K 0.0)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   187
val e_fun_weight_span =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   188
  Attrib.setup_config_real @{binding atp_e_fun_weight_span} (K 40.0)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   189
val e_default_sym_offs_weight =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   190
  Attrib.setup_config_real @{binding atp_e_default_sym_offs_weight} (K 1.0)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   191
val e_sym_offs_weight_base =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   192
  Attrib.setup_config_real @{binding atp_e_sym_offs_weight_base} (K ~20.0)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   193
val e_sym_offs_weight_span =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   194
  Attrib.setup_config_real @{binding atp_e_sym_offs_weight_span} (K 60.0)
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   195
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42332
diff changeset
   196
fun e_weight_method_case method fw sow =
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   197
  if method = e_fun_weightN then fw
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   198
  else if method = e_sym_offset_weightN then sow
43478
1cef683b588b slightly better setup for E
blanchet
parents: 43476
diff changeset
   199
  else raise Fail ("unexpected " ^ quote method)
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   200
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   201
fun scaled_e_weight ctxt method w =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   202
  w * Config.get ctxt
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   203
          (e_weight_method_case method e_fun_weight_span e_sym_offs_weight_span)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   204
  + Config.get ctxt
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   205
        (e_weight_method_case method 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
   206
  |> 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
   207
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   208
fun e_weight_arguments ctxt method weights =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   209
  if method = e_autoN then
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   210
    "-xAutoDev"
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   211
  else
43622
blanchet
parents: 43575
diff changeset
   212
    (* supplied by Stephan Schulz *)
41314
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   213
    "--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
   214
    \--destructive-er-aggressive --destructive-er --presat-simplify \
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   215
    \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   216
    \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42332
diff changeset
   217
    \-H'(4*" ^ e_weight_method_case method "FunWeight" "SymOffsetWeight" ^
41725
7cca2de89296 added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents: 41335
diff changeset
   218
    "(SimulateSOS, " ^
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   219
    (e_weight_method_case method e_default_fun_weight e_default_sym_offs_weight
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   220
     |> 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
   221
    ",20,1.5,1.5,1" ^
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   222
    (weights ()
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   223
     |> map (fn (s, w) => "," ^ s ^ ":" ^ scaled_e_weight ctxt method w)
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   224
     |> implode) ^
41314
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   225
    "),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
   226
    \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\
2dc1dfc1bc69 use the options provided by Stephan Schulz -- much better
blanchet
parents: 41313
diff changeset
   227
    \FIFOWeight(PreferProcessed))'"
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41269
diff changeset
   228
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   229
fun effective_e_weight_method ctxt =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   230
  if is_old_e_version () then e_autoN else Config.get ctxt e_weight_method
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42332
diff changeset
   231
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   232
val e_config : atp_config =
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   233
  {exec = ("E_HOME", "eproof"),
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   234
   required_execs = [],
43354
396aaa15dd8b pass --trim option to "eproof" script to speed up proof reconstruction
blanchet
parents: 43288
diff changeset
   235
   arguments =
43567
dda3e38cc351 remove experimental trimming feature -- it slowed down things on Linux for some reason
blanchet
parents: 43566
diff changeset
   236
     fn ctxt => fn _ => fn method => fn timeout => fn weights =>
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   237
        "--tstp-in --tstp-out -l5 " ^ e_weight_arguments ctxt method weights ^
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
   238
        " -tAutoDev --silent --cpu-limit=" ^ string_of_int (to_secs 2 timeout),
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42955
diff changeset
   239
   proof_delims = tstp_proof_delims,
36265
41c9e755e552 distinguish between the different ATP errors in the user interface;
blanchet
parents: 36264
diff changeset
   240
   known_failures =
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   241
     known_szs_status_failures @
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   242
     [(TimedOut, "Failure: Resource limit exceeded (time)"),
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   243
      (TimedOut, "time limit exceeded"),
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   244
      (OutOfResources, "# Cannot determine problem status")],
43466
52f040bcfae7 tweaked TPTP formula kind for typing information used in the conjecture
blanchet
parents: 43354
diff changeset
   245
   conj_sym_kind = Hypothesis,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   246
   prem_kind = Conjecture,
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42643
diff changeset
   247
   best_slices = fn ctxt =>
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   248
     let val method = effective_e_weight_method ctxt in
43474
423cd1ecf714 optimized E's time slicing, based on latest exhaustive Judgment Day results
blanchet
parents: 43473
diff changeset
   249
       (* FUDGE *)
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   250
       if method = e_smartN then
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   251
         [(0.333, (true, ((500, FOF, "mono_tags??", combsN), e_fun_weightN))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   252
          (0.334, (true, ((50, FOF, "mono_guards??", combsN), e_fun_weightN))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   253
          (0.333, (true, ((1000, FOF, "mono_tags??", combsN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   254
                          e_sym_offset_weightN)))]
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   255
       else
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   256
         [(1.0, (true, ((500, FOF, "mono_tags??", combsN), method)))]
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   257
     end}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   258
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   259
val e = (eN, e_config)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   260
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   261
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   262
(* LEO-II *)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   263
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   264
val leo2_thf0 = THF (TPTP_Monomorphic, TPTP_Explicit, THF_Without_Choice)
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   265
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   266
val leo2_config : atp_config =
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   267
  {exec = ("LEO2_HOME", "leo"),
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   268
   required_execs = [],
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   269
   arguments =
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   270
     fn _ => fn _ => fn sos => fn timeout => fn _ =>
45300
d8c8c2fcab2c specify proof output level 1 (i.e. no detailed, potentially huge E proofs) to LEO-II; requires version 1.2.9
blanchet
parents: 45234
diff changeset
   271
        "--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   272
        |> sos = sosN ? prefix "--sos ",
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   273
   proof_delims = tstp_proof_delims,
45207
1d13334628a9 one more LEO-II failure
blanchet
parents: 45203
diff changeset
   274
   known_failures =
1d13334628a9 one more LEO-II failure
blanchet
parents: 45203
diff changeset
   275
     known_szs_status_failures @
1d13334628a9 one more LEO-II failure
blanchet
parents: 45203
diff changeset
   276
     [(TimedOut, "CPU time limit exceeded, terminating")],
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   277
   conj_sym_kind = Axiom,
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   278
   prem_kind = Hypothesis,
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   279
   best_slices = fn ctxt =>
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   280
     (* FUDGE *)
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   281
     [(0.667, (false, ((150, leo2_thf0, "mono_simple_higher", liftingN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   282
                       sosN))),
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   283
      (0.333, (true, ((50, leo2_thf0, "mono_simple_higher", liftingN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   284
                      no_sosN)))]
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   285
     |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   286
         else I)}
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   287
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   288
val leo2 = (leo2N, leo2_config)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   289
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   290
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   291
(* Satallax *)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   292
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   293
val satallax_thf0 = THF (TPTP_Monomorphic, TPTP_Explicit, THF_With_Choice)
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   294
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   295
val satallax_config : atp_config =
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   296
  {exec = ("SATALLAX_HOME", "satallax"),
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   297
   required_execs = [],
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   298
   arguments =
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   299
     fn _ => fn _ => fn _ => fn timeout => fn _ =>
45162
170dffc6df75 parse Satallax unsat cores
blanchet
parents: 44786
diff changeset
   300
        "-p hocore -t " ^ string_of_int (to_secs 1 timeout),
170dffc6df75 parse Satallax unsat cores
blanchet
parents: 44786
diff changeset
   301
   proof_delims =
170dffc6df75 parse Satallax unsat cores
blanchet
parents: 44786
diff changeset
   302
     [("% Higher-Order Unsat Core BEGIN", "% Higher-Order Unsat Core END")],
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   303
   known_failures = known_szs_status_failures,
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   304
   conj_sym_kind = Axiom,
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   305
   prem_kind = Hypothesis,
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   306
   best_slices =
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   307
     (* FUDGE *)
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   308
     K [(1.0, (true, ((100, satallax_thf0, "mono_simple_higher", keep_lamsN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   309
                      "")))]}
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   310
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   311
val satallax = (satallaxN, satallax_config)
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   312
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   313
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   314
(* 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
   315
46380
7e049e9f5c8b new SPASS setup
blanchet
parents: 46370
diff changeset
   316
val spass_incompleteN = "incomplete"
7e049e9f5c8b new SPASS setup
blanchet
parents: 46370
diff changeset
   317
36219
16670b4f0baa set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents: 36190
diff changeset
   318
(* The "-VarWeight=3" option helps the higher-order problems, probably by
44450
d848dd7b21f4 fixed "hBOOL" of existential variables, and generate more helpers
blanchet
parents: 44425
diff changeset
   319
   counteracting the presence of explicit application operators. *)
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   320
val spass_config : atp_config =
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   321
  {exec = ("ISABELLE_ATP", "scripts/spass"),
39002
a2d7be688ea1 add dependency of "spass" script
blanchet
parents: 38999
diff changeset
   322
   required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")],
43569
b342cd125533 removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents: 43567
diff changeset
   323
   arguments = fn _ => fn _ => fn sos => fn timeout => fn _ =>
37962
d7dbe01f48d7 keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents: 37926
diff changeset
   324
     ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
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
   325
      \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 1 timeout))
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   326
     |> sos = sosN ? prefix "-SOS=1 ",
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   327
   proof_delims = [("Here is a proof", "Formulae used in the proof")],
36289
f75b6a3e1450 set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents: 36287
diff changeset
   328
   known_failures =
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   329
     known_perl_failures @
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 42999
diff changeset
   330
     [(GaveUp, "SPASS beiseite: Completion found"),
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   331
      (TimedOut, "SPASS beiseite: Ran out of time"),
36965
67ae217c6b5c identify common SPASS error more clearly
blanchet
parents: 36924
diff changeset
   332
      (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
37413
e856582fe9c4 improve ATP-specific error messages
blanchet
parents: 37347
diff changeset
   333
      (MalformedInput, "Undefined symbol"),
37414
d0cea0796295 expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet
parents: 37413
diff changeset
   334
      (MalformedInput, "Free Variable"),
44391
7b4104df2be6 gracefully handle empty SPASS problems
blanchet
parents: 44235
diff changeset
   335
      (Unprovable, "No formulae and clauses found in input file"),
39263
e2a3c435334b more precise error messages when Vampire is interrupted or SPASS runs into an internal bug
blanchet
parents: 39262
diff changeset
   336
      (InternalError, "Please report this error")],
43466
52f040bcfae7 tweaked TPTP formula kind for typing information used in the conjecture
blanchet
parents: 43354
diff changeset
   337
   conj_sym_kind = Hypothesis,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   338
   prem_kind = Conjecture,
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
   339
   best_slices = fn ctxt =>
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
   340
     (* FUDGE *)
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   341
     [(0.333, (false, ((150, DFG DFG_Unsorted, "mono_tags??", liftingN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   342
                       sosN))),
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   343
      (0.333, (false, ((300, DFG DFG_Unsorted, "poly_tags??", liftingN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   344
                       sosN))),
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   345
      (0.334, (false, ((50, DFG DFG_Unsorted, "mono_tags??", liftingN),
45876
40952db4e57b SPASS is incomplete because of the -Splits and -FullRed options, not just because of -SOS=1 -- don't pretend the opposite
blanchet
parents: 45521
diff changeset
   346
                       no_sosN)))]
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   347
     |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single
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
   348
         else I)}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   349
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   350
val spass = (spassN, spass_config)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   351
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   352
val spass_new_macro_slice_1 = (300, DFG DFG_Sorted, "mono_simple", combsN)
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   353
val spass_new_macro_slice_2 = (50, DFG DFG_Sorted, "mono_simple", combsN)
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   354
val spass_new_macro_slice_3 = (150, DFG DFG_Sorted, "mono_simple", liftingN)
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   355
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   356
(* Experimental *)
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   357
val spass_new_config : atp_config =
46370
b3e53dd6f10a new SPASS setup
blanchet
parents: 46365
diff changeset
   358
  {exec = ("ISABELLE_ATP", "scripts/spass_new"),
46398
caf27e675dd1 include new SPASS by default if available
blanchet
parents: 46385
diff changeset
   359
   required_execs =
caf27e675dd1 include new SPASS by default if available
blanchet
parents: 46385
diff changeset
   360
     [("SPASS_NEW_HOME", "SPASS"), ("SPASS_NEW_HOME", "tptp2dfg")],
46380
7e049e9f5c8b new SPASS setup
blanchet
parents: 46370
diff changeset
   361
   arguments = fn _ => fn _ => fn incomplete => fn timeout => fn _ =>
46402
ef8d65f64f77 change 9ce354a77908 wasn't quite right -- here's an improvement
blanchet
parents: 46401
diff changeset
   362
     ("-Auto -LR=1 -LT=1 -Isabelle=1 -TimeLimit=" ^
ef8d65f64f77 change 9ce354a77908 wasn't quite right -- here's an improvement
blanchet
parents: 46401
diff changeset
   363
      string_of_int (to_secs 1 timeout))
46380
7e049e9f5c8b new SPASS setup
blanchet
parents: 46370
diff changeset
   364
     |> incomplete = spass_incompleteN ? prefix "-Splits=0 -FullRed=0 ",
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   365
   proof_delims = #proof_delims spass_config,
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   366
   known_failures = #known_failures spass_config,
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   367
   conj_sym_kind = #conj_sym_kind spass_config,
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   368
   prem_kind = #prem_kind spass_config,
46381
ef62c2fafa9e fixed syntax bug in DFG output
blanchet
parents: 46380
diff changeset
   369
   best_slices = fn _ =>
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   370
     (* FUDGE *)
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   371
     [(0.300, (true, (spass_new_macro_slice_1, ""))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   372
      (0.333, (true, (spass_new_macro_slice_2, ""))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   373
      (0.333, (true, (spass_new_macro_slice_3, "")))]}
45301
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   374
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   375
val spass_new = (spass_newN, spass_new_config)
866b075aa99b added sorted DFG output for coming version of SPASS
blanchet
parents: 45300
diff changeset
   376
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   377
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   378
(* Vampire *)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   379
44507
e08158671ef4 disable TFF for Vampire 1.8 until they've fixed the soundness issues and it's back on SystemOnTPTP
blanchet
parents: 44503
diff changeset
   380
(* Vampire 1.8 has TFF support, but it's buggy and therefore disabled on
e08158671ef4 disable TFF for Vampire 1.8 until they've fixed the soundness issues and it's back on SystemOnTPTP
blanchet
parents: 44503
diff changeset
   381
   SystemOnTPTP. *)
44420
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   382
fun is_old_vampire_version () =
44507
e08158671ef4 disable TFF for Vampire 1.8 until they've fixed the soundness issues and it's back on SystemOnTPTP
blanchet
parents: 44503
diff changeset
   383
  string_ord (getenv "VAMPIRE_VERSION", "1.8") <> GREATER
44420
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   384
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   385
val vampire_tff0 = TFF (TPTP_Monomorphic, TPTP_Implicit)
44589
0a1dfc6365e9 first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents: 44586
diff changeset
   386
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   387
val vampire_config : atp_config =
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   388
  {exec = ("VAMPIRE_HOME", "vampire"),
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   389
   required_execs = [],
43569
b342cd125533 removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents: 43567
diff changeset
   390
   arguments = fn _ => fn _ => fn sos => fn timeout => fn _ =>
44417
c76c04d876ef kindly ask Vampire to output axiom names
blanchet
parents: 44416
diff changeset
   391
     "--mode casc -t " ^ string_of_int (to_secs 1 timeout) ^
45234
5509362b924b disable Vampire's BDD optimization, which sometimes yields so huge proofs that this causes problems for reconstruction
blanchet
parents: 45207
diff changeset
   392
     " --proof tptp --output_axiom_names on\
5509362b924b disable Vampire's BDD optimization, which sometimes yields so huge proofs that this causes problems for reconstruction
blanchet
parents: 45207
diff changeset
   393
     \ --forced_options propositional_to_bdd=off\
44417
c76c04d876ef kindly ask Vampire to output axiom names
blanchet
parents: 44416
diff changeset
   394
     \ --thanks \"Andrei and Krystof\" --input_file"
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   395
     |> sos = sosN ? prefix "--sos on ",
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   396
   proof_delims =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   397
     [("=========== Refutation ==========",
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   398
       "======= End of refutation ======="),
38033
df99f022751d support latest version of Vampire (1.0) locally
blanchet
parents: 38032
diff changeset
   399
      ("% SZS output start Refutation", "% SZS output end Refutation"),
df99f022751d support latest version of Vampire (1.0) locally
blanchet
parents: 38032
diff changeset
   400
      ("% SZS output start Proof", "% SZS output end Proof")],
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   401
   known_failures =
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   402
     known_szs_status_failures @
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 42999
diff changeset
   403
     [(GaveUp, "UNPROVABLE"),
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 42999
diff changeset
   404
      (GaveUp, "CANNOT PROVE"),
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   405
      (Unprovable, "Satisfiability detected"),
38647
5500241da479 play with fudge factor + parse one more Vampire error
blanchet
parents: 38646
diff changeset
   406
      (Unprovable, "Termination reason: Satisfiable"),
39263
e2a3c435334b more precise error messages when Vampire is interrupted or SPASS runs into an internal bug
blanchet
parents: 39262
diff changeset
   407
      (Interrupted, "Aborted by signal SIGINT")],
43466
52f040bcfae7 tweaked TPTP formula kind for typing information used in the conjecture
blanchet
parents: 43354
diff changeset
   408
   conj_sym_kind = Conjecture,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   409
   prem_kind = Conjecture,
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
   410
   best_slices = fn ctxt =>
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
   411
     (* FUDGE *)
44420
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   412
     (if is_old_vampire_version () then
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   413
        [(0.333, (false, ((150, FOF, "poly_guards??", combs_or_liftingN),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   414
                           sosN))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   415
         (0.333, (false, ((500, FOF, "mono_tags??", combs_or_liftingN), sosN))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   416
         (0.334, (true, ((50, FOF, "mono_guards??", combs_or_liftingN),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   417
                         no_sosN)))]
44420
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   418
      else
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   419
        [(0.333, (false, ((150, vampire_tff0, "poly_guards??",
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   420
                           combs_or_liftingN), sosN))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   421
         (0.333, (false, ((500, vampire_tff0, "mono_simple", combs_or_liftingN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   422
                          sosN))),
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   423
         (0.334, (true, ((50, vampire_tff0, "mono_simple", combs_or_liftingN),
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   424
                         no_sosN)))])
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   425
     |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single
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
   426
         else I)}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   427
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   428
val vampire = (vampireN, vampire_config)
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   429
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   430
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   431
(* Z3 with TPTP syntax *)
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   432
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   433
val z3_tff0 = TFF (TPTP_Monomorphic, TPTP_Implicit)
44589
0a1dfc6365e9 first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents: 44586
diff changeset
   434
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
   435
val z3_tptp_config : atp_config =
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   436
  {exec = ("Z3_HOME", "z3"),
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   437
   required_execs = [],
43354
396aaa15dd8b pass --trim option to "eproof" script to speed up proof reconstruction
blanchet
parents: 43288
diff changeset
   438
   arguments = fn _ => fn _ => fn _ => fn timeout => fn _ =>
44420
3d0921b91a86 avoid TFF format with older Vampire versions
blanchet
parents: 44418
diff changeset
   439
     "MBQI=true -tptp -t:" ^ string_of_int (to_secs 1 timeout),
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   440
   proof_delims = [],
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   441
   known_failures = known_szs_status_failures,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   442
   conj_sym_kind = Hypothesis,
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   443
   prem_kind = Hypothesis,
42723
c1909691bbf0 allow each slice to have its own type system
blanchet
parents: 42710
diff changeset
   444
   best_slices =
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
   445
     (* FUDGE *)
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   446
     K [(0.5, (false, ((250, z3_tff0, "mono_simple", combsN), ""))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   447
        (0.25, (false, ((125, z3_tff0, "mono_simple", combsN), ""))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   448
        (0.125, (false, ((62, z3_tff0, "mono_simple", combsN), ""))),
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   449
        (0.125, (false, ((31, z3_tff0, "mono_simple", combsN), "")))]}
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   450
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
   451
val z3_tptp = (z3_tptpN, z3_tptp_config)
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   452
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   453
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   454
(* Not really a prover: Experimental Polymorphic TFF and THF output *)
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   455
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   456
fun dummy_config format type_enc : atp_config =
44596
2621046c550a cleaner "pff" dummy TFF0 prover
blanchet
parents: 44591
diff changeset
   457
  {exec = ("ISABELLE_ATP", "scripts/dummy_atp"),
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   458
   required_execs = [],
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   459
   arguments = K (K (K (K (K "")))),
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   460
   proof_delims = [],
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   461
   known_failures = known_szs_status_failures,
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   462
   conj_sym_kind = Hypothesis,
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   463
   prem_kind = Hypothesis,
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   464
   best_slices =
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   465
     K [(1.0, (false, ((200, format, type_enc,
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   466
                        if is_format_higher_order format then keep_lamsN
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   467
                        else combsN), "")))]}
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   468
45365
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   469
val dummy_tff1_format = TFF (TPTP_Polymorphic, TPTP_Explicit)
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   470
val dummy_tff1_config = dummy_config dummy_tff1_format "poly_simple"
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   471
val dummy_tff1 = (dummy_tff1N, dummy_tff1_config)
44590
3a8a31be92d2 added dummy PFF prover, for debugging purposes
blanchet
parents: 44589
diff changeset
   472
45365
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   473
val dummy_thf_format = THF (TPTP_Polymorphic, TPTP_Explicit, THF_With_Choice)
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   474
val dummy_thf_config = dummy_config dummy_thf_format "poly_simple_higher"
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   475
val dummy_thf = (dummy_thfN, dummy_thf_config)
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   476
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   477
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   478
(* Remote ATP invocation via SystemOnTPTP *)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   479
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   480
val systems = Synchronized.var "atp_systems" ([] : string list)
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   481
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   482
fun get_systems () =
44589
0a1dfc6365e9 first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents: 44586
diff changeset
   483
  case Isabelle_System.bash_output
0a1dfc6365e9 first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents: 44586
diff changeset
   484
           "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   485
    (output, 0) => split_lines output
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   486
  | (output, _) =>
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   487
    error (case extract_known_failure known_perl_failures output of
41744
a18e7bbca258 make minimizer verbose
blanchet
parents: 41742
diff changeset
   488
             SOME failure => string_for_failure failure
39491
2416666e6f94 refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents: 39375
diff changeset
   489
           | NONE => perhaps (try (unsuffix "\n")) output ^ ".")
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   490
42537
25ceb855a18b improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
blanchet
parents: 42535
diff changeset
   491
fun find_system name [] systems =
25ceb855a18b improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
blanchet
parents: 42535
diff changeset
   492
    find_first (String.isPrefix (name ^ "---")) 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
   493
  | find_system name (version :: versions) systems =
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   494
    case find_first (String.isPrefix (name ^ "---" ^ version)) systems of
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   495
      NONE => find_system name versions systems
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   496
    | 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
   497
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   498
fun get_system name versions =
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   499
  Synchronized.change_result systems
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   500
      (fn systems => (if null systems then get_systems () else systems)
42955
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   501
                     |> `(`(find_system name versions)))
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
   502
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
   503
fun the_system name versions =
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   504
  case get_system name versions of
42955
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   505
    (SOME sys, _) => sys
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   506
  | (NONE, []) => error ("SystemOnTPTP is currently not available.")
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   507
  | (NONE, syss) =>
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   508
    error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n" ^
576bd30cc4ea clearer SystemOnTPTP errors
blanchet
parents: 42954
diff changeset
   509
           "(Available systems: " ^ commas_quote syss ^ ".)")
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   510
41148
f5229ab54284 added timeout max for remote server invocation
blanchet
parents: 40669
diff changeset
   511
val max_remote_secs = 240 (* give Geoff Sutcliffe's servers a break *)
f5229ab54284 added timeout max for remote server invocation
blanchet
parents: 40669
diff changeset
   512
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
   513
fun remote_config system_name system_versions proof_delims known_failures
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   514
                  conj_sym_kind prem_kind best_slice : atp_config =
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   515
  {exec = ("ISABELLE_ATP", "scripts/remote_atp"),
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38090
diff changeset
   516
   required_execs = [],
43354
396aaa15dd8b pass --trim option to "eproof" script to speed up proof reconstruction
blanchet
parents: 43288
diff changeset
   517
   arguments = fn _ => fn _ => fn _ => fn timeout => fn _ =>
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
   518
     "-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout))
41148
f5229ab54284 added timeout max for remote server invocation
blanchet
parents: 40669
diff changeset
   519
     ^ " -s " ^ the_system system_name system_versions,
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42955
diff changeset
   520
   proof_delims = union (op =) tstp_proof_delims proof_delims,
45203
e3c13fa443ef more uniform SZS status handling
blanchet
parents: 45162
diff changeset
   521
   known_failures = known_failures @ known_perl_failures @ known_says_failures,
42709
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   522
   conj_sym_kind = conj_sym_kind,
e7af132d48fe allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents: 42708
diff changeset
   523
   prem_kind = prem_kind,
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   524
   best_slices = fn ctxt =>
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   525
     let val (max_relevant, format, type_enc, lam_trans) = best_slice ctxt in
46407
30e9720cc0b9 optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents: 46402
diff changeset
   526
       [(1.0, (false, ((max_relevant, format, type_enc, lam_trans), "")))]
43473
fb2713b803e6 deal with ATP time slices in a more flexible/robust fashion
blanchet
parents: 43467
diff changeset
   527
     end}
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42332
diff changeset
   528
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   529
fun remotify_config system_name system_versions best_slice
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   530
        ({proof_delims, known_failures, conj_sym_kind, prem_kind, ...}
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   531
         : atp_config) : atp_config =
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
   532
  remote_config system_name system_versions proof_delims known_failures
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   533
                conj_sym_kind prem_kind best_slice
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   534
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   535
fun remote_atp name system_name system_versions proof_delims known_failures
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   536
               conj_sym_kind prem_kind best_slice =
40060
5ef6747aa619 first step in adding support for an SMT backend to Sledgehammer
blanchet
parents: 40059
diff changeset
   537
  (remote_prefix ^ name,
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
   538
   remote_config system_name system_versions proof_delims known_failures
44416
cabd06b69c18 added formats to the slice and use TFF for remote Vampire
blanchet
parents: 44391
diff changeset
   539
                 conj_sym_kind prem_kind best_slice)
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   540
fun remotify_atp (name, config) system_name system_versions best_slice =
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   541
  (remote_prefix ^ name,
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   542
   remotify_config system_name system_versions best_slice config)
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
   543
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   544
val explicit_tff0 = TFF (TPTP_Monomorphic, TPTP_Explicit)
44589
0a1dfc6365e9 first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents: 44586
diff changeset
   545
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   546
val remote_e =
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   547
  remotify_atp e "EP" ["1.0", "1.1", "1.2"]
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   548
      (K (750, FOF, "mono_tags??", combsN) (* FUDGE *))
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   549
val remote_leo2 =
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   550
  remotify_atp leo2 "LEO-II" ["1.2.8", "1.2.6"]
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   551
      (K (100, leo2_thf0, "mono_simple_higher", liftingN) (* FUDGE *))
44099
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   552
val remote_satallax =
5e14f591515e support local HOATPs
blanchet
parents: 44096
diff changeset
   553
  remotify_atp satallax "Satallax" ["2.1", "2.0", "2"]
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   554
      (K (100, satallax_thf0, "mono_simple_higher", keep_lamsN) (* FUDGE *))
43500
4c357b7aa710 provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents: 43497
diff changeset
   555
val remote_vampire =
44499
8870232a87ad make TFF output less explicit where possible
blanchet
parents: 44498
diff changeset
   556
  remotify_atp vampire "Vampire" ["1.8"]
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   557
      (K (250, FOF, "mono_guards??", combs_or_liftingN) (* FUDGE *))
44423
f74707e12d30 exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents: 44422
diff changeset
   558
val remote_z3_tptp =
44754
265174356212 added dummy polymorphic THF system
blanchet
parents: 44596
diff changeset
   559
  remotify_atp z3_tptp "Z3" ["3.0"]
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   560
      (K (250, z3_tff0, "mono_simple", combsN) (* FUDGE *))
44092
bf489e54d7f8 renamed E wrappers for consistency with CASC conventions
blanchet
parents: 43989
diff changeset
   561
val remote_e_sine =
bf489e54d7f8 renamed E wrappers for consistency with CASC conventions
blanchet
parents: 43989
diff changeset
   562
  remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Axiom
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   563
      Conjecture (K (500, FOF, "mono_guards??", combsN) (* FUDGE *))
45338
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
   564
val remote_iprover =
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
   565
  remote_atp iproverN "iProver" [] [] [] Axiom Conjecture
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   566
      (K (150, FOF, "mono_guards??", liftingN) (* FUDGE *))
45338
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
   567
val remote_iprover_eq =
b9d5d3625e9a added remote iProver(-Eq) for experimentation
blanchet
parents: 45304
diff changeset
   568
  remote_atp iprover_eqN "iProver-Eq" [] [] [] Axiom Conjecture
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   569
      (K (150, FOF, "mono_guards??", liftingN) (* FUDGE *))
41740
4b09f8b9e012 added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents: 41738
diff changeset
   570
val remote_snark =
42939
0134d6650092 added support for remote Waldmeister
blanchet
parents: 42938
diff changeset
   571
  remote_atp snarkN "SNARK" ["20080805r029", "20080805r024"]
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   572
      [("refutation.", "end_refutation.")] [] Hypothesis Hypothesis
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   573
      (K (100, explicit_tff0, "mono_simple", liftingN) (* FUDGE *))
44092
bf489e54d7f8 renamed E wrappers for consistency with CASC conventions
blanchet
parents: 43989
diff changeset
   574
val remote_e_tofof =
44589
0a1dfc6365e9 first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents: 44586
diff changeset
   575
  remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config) Axiom
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   576
      Hypothesis
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   577
      (K (150, explicit_tff0, "mono_simple", liftingN) (* FUDGE *))
42938
c124032ac96f added Waldmeister
blanchet
parents: 42937
diff changeset
   578
val remote_waldmeister =
c124032ac96f added Waldmeister
blanchet
parents: 42937
diff changeset
   579
  remote_atp waldmeisterN "Waldmeister" ["710"]
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   580
      [("#START OF PROOF", "Proved Goals:")]
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   581
      [(OutOfResources, "Too many function symbols"),
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   582
       (Crashed, "Unrecoverable Segmentation Fault")]
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45365
diff changeset
   583
      Hypothesis Hypothesis
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46320
diff changeset
   584
      (K (50, CNF_UEQ, "mono_tags??", combsN) (* FUDGE *))
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   585
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   586
(* Setup *)
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   587
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   588
fun add_atp (name, config) thy =
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   589
  Data.map (Symtab.update_new (name, (config, stamp ()))) thy
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   590
  handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   591
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   592
fun get_atp thy name =
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   593
  the (Symtab.lookup (Data.get thy) name) |> fst
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   594
  handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   595
41727
ab3f6d76fb23 available_provers ~> supported_provers (for clarity)
blanchet
parents: 41725
diff changeset
   596
val supported_atps = Symtab.keys o Data.get
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   597
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   598
fun is_atp_installed thy name =
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   599
  let val {exec, required_execs, ...} = get_atp thy name in
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   600
    forall (curry (op <>) "" o getenv o fst) (exec :: required_execs)
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   601
  end
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   602
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   603
fun refresh_systems_on_tptp () =
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   604
  Synchronized.change systems (fn _ => get_systems ())
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   605
42962
3b50fdeb6cfc started adding support for THF output (but no lambdas)
blanchet
parents: 42955
diff changeset
   606
val atps =
45365
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   607
  [e, leo2, dummy_tff1, dummy_thf, satallax, spass, spass_new, vampire, z3_tptp,
c71e6980ad28 renamed experimental systems
blanchet
parents: 45339
diff changeset
   608
   remote_e, remote_e_sine, remote_e_tofof, remote_iprover, remote_iprover_eq,
45339
4f6ae5423311 document new experimental provers
blanchet
parents: 45338
diff changeset
   609
   remote_leo2, remote_satallax, remote_vampire, remote_z3_tptp, remote_snark,
4f6ae5423311 document new experimental provers
blanchet
parents: 45338
diff changeset
   610
   remote_waldmeister]
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39491
diff changeset
   611
val setup = fold add_atp atps
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
   612
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
   613
end;