src/HOL/Tools/ATP/atp_systems.ML
author blanchet
Tue, 24 Aug 2010 14:36:29 +0200
changeset 38690 38a926e033ad
parent 38685 87a1e97a3ef3
child 38691 fe5929dacd43
permissions -rw-r--r--
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
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
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    10
  datatype failure =
38519
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
    11
    Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
    12
    OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
    13
    MalformedInput | MalformedOutput | UnknownError
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    14
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    15
  type prover_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
    16
    {exec: string * string,
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
    17
     required_execs: (string * string) list,
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    18
     arguments: bool -> Time.time -> string,
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
    19
     has_incomplete_mode: bool,
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    20
     proof_delims: (string * string) list,
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    21
     known_failures: (failure * string) list,
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
    22
     default_max_relevant_per_iter: int,
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
    23
     default_theory_relevant: bool,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
    24
     explicit_forall: bool,
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
    25
     use_conjecture_for_hypotheses: bool}
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    26
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    27
  val string_for_failure : failure -> string
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    28
  val known_failure_in_output :
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    29
    string -> (failure * string) list -> failure option
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    30
  val add_prover: string * prover_config -> theory -> theory
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    31
  val get_prover: theory -> string -> prover_config
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    32
  val available_atps: theory -> unit
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
    33
  val refresh_systems_on_tptp : unit -> unit
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
    34
  val default_atps_param_value : unit -> string
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
    35
  val setup : theory -> theory
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    36
end;
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    37
36376
e83d52a52449 renamed module "ATP_Wrapper" to "ATP_Systems"
blanchet
parents: 36371
diff changeset
    38
structure ATP_Systems : ATP_SYSTEMS =
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
    39
struct
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
    40
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    41
(* prover configuration *)
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
    42
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
    43
datatype failure =
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
    44
  Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
38519
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
    45
  SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
    46
  MalformedOutput | UnknownError
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
    47
32941
72d48e333b77 eliminated extraneous wrapping of public records;
wenzelm
parents: 32936
diff changeset
    48
type prover_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
    49
  {exec: string * string,
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
    50
   required_execs: (string * string) list,
37514
b147d01b8ebc if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents: 37509
diff changeset
    51
   arguments: bool -> Time.time -> string,
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
    52
   has_incomplete_mode: bool,
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    53
   proof_delims: (string * string) list,
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    54
   known_failures: (failure * string) list,
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
    55
   default_max_relevant_per_iter: int,
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
    56
   default_theory_relevant: bool,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
    57
   explicit_forall: bool,
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
    58
   use_conjecture_for_hypotheses: bool}
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
    59
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    60
val missing_message_tail =
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    61
  " appears to be missing. You will need to install it if you want to run \
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    62
  \ATPs remotely."
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    63
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    64
fun string_for_failure Unprovable = "The ATP problem is unprovable."
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    65
  | string_for_failure IncompleteUnprovable =
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    66
    "The ATP cannot prove the problem."
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38092
diff changeset
    67
  | string_for_failure CantConnect = "Can't connect to remote server."
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    68
  | string_for_failure TimedOut = "Timed out."
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    69
  | string_for_failure OutOfResources = "The ATP ran out of resources."
38519
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
    70
  | string_for_failure SpassTooOld =
38096
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    71
    "Isabelle requires a more recent version of SPASS with support for the \
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    72
    \TPTP syntax. To install it, download and extract the package \
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    73
    \\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    74
    \\"spass-3.7\" directory's absolute path to " ^
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    75
    quote (Path.implode (Path.expand (Path.appends
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    76
               (Path.variable "ISABELLE_HOME_USER" ::
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    77
                map Path.basic ["etc", "components"])))) ^
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    78
    " on a line of its own."
38519
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
    79
  | string_for_failure VampireTooOld =
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
    80
    "Isabelle requires a more recent version of Vampire. To install it, follow \
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
    81
    \the instructions from the Sledgehammer manual (\"isabelle doc\
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
    82
    \ sledgehammer\")."
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    83
  | string_for_failure NoPerl = "Perl" ^ missing_message_tail
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    84
  | string_for_failure NoLibwwwPerl =
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    85
    "The Perl module \"libwww-perl\"" ^ missing_message_tail
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    86
  | string_for_failure MalformedInput =
38096
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    87
    "The ATP problem is malformed. Please report this to the Isabelle \
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    88
    \developers."
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    89
  | string_for_failure MalformedOutput = "The ATP output is malformed."
488b38cd3e06 avoid "ATP Error: Error: blah" style messages
blanchet
parents: 38094
diff changeset
    90
  | string_for_failure UnknownError = "An unknown ATP error occurred."
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    91
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    92
fun known_failure_in_output output =
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    93
  find_first (fn (_, pattern) => String.isSubstring pattern output)
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    94
  #> Option.map fst
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    95
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
    96
val known_perl_failures =
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38092
diff changeset
    97
  [(CantConnect, "HTTP error"),
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38092
diff changeset
    98
   (NoPerl, "env: perl"),
38065
9069e1ad1527 improved ATP error handling some more
blanchet
parents: 38064
diff changeset
    99
   (NoLibwwwPerl, "Can't locate HTTP")]
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   100
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   101
(* named provers *)
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   102
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   103
structure Data = Theory_Data
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   104
(
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   105
  type T = (prover_config * stamp) Symtab.table
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   106
  val empty = Symtab.empty
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   107
  val extend = I
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   108
  fun merge data : T = Symtab.merge (eq_snd op =) data
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   109
    handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   110
)
38017
3ad3e3ca2451 move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents: 38015
diff changeset
   111
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   112
fun add_prover (name, config) thy =
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   113
  Data.map (Symtab.update_new (name, (config, stamp ()))) thy
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   114
  handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
38017
3ad3e3ca2451 move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents: 38015
diff changeset
   115
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   116
fun get_prover thy name =
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   117
  the (Symtab.lookup (Data.get thy) name) |> fst
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   118
  handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
37962
d7dbe01f48d7 keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents: 37926
diff changeset
   119
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   120
fun available_atps thy =
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   121
  priority ("Available ATPs: " ^
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   122
            commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32740
diff changeset
   123
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   124
fun available_atps thy =
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   125
  priority ("Available ATPs: " ^
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   126
            commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   127
36382
b90fc0d75bca cosmetics
blanchet
parents: 36377
diff changeset
   128
fun to_generous_secs time = (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
   129
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   130
(* E prover *)
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   131
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   132
val tstp_proof_delims =
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   133
  ("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation")
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   134
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35869
diff changeset
   135
val e_config : prover_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
   136
  {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
   137
   required_execs = [],
37514
b147d01b8ebc if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents: 37509
diff changeset
   138
   arguments = fn _ => fn timeout =>
36382
b90fc0d75bca cosmetics
blanchet
parents: 36377
diff changeset
   139
     "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^
b90fc0d75bca cosmetics
blanchet
parents: 36377
diff changeset
   140
     string_of_int (to_generous_secs timeout),
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
   141
   has_incomplete_mode = false,
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   142
   proof_delims = [tstp_proof_delims],
36265
41c9e755e552 distinguish between the different ATP errors in the user interface;
blanchet
parents: 36264
diff changeset
   143
   known_failures =
37995
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37994
diff changeset
   144
     [(Unprovable, "SZS status: CounterSatisfiable"),
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37994
diff changeset
   145
      (Unprovable, "SZS status CounterSatisfiable"),
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   146
      (TimedOut, "Failure: Resource limit exceeded (time)"),
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   147
      (TimedOut, "time limit exceeded"),
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   148
      (OutOfResources,
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   149
       "# Cannot determine problem status within resource limit"),
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   150
      (OutOfResources, "SZS status: ResourceOut"),
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   151
      (OutOfResources, "SZS status ResourceOut")],
38685
87a1e97a3ef3 adjust fudge factors in the light of experiments
blanchet
parents: 38680
diff changeset
   152
   default_max_relevant_per_iter = 80 (* FUDGE *),
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   153
   default_theory_relevant = false,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   154
   explicit_forall = false,
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   155
   use_conjecture_for_hypotheses = true}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   156
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   157
val e = ("e", e_config)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   158
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   159
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
   160
(* The "-VarWeight=3" option helps the higher-order problems, probably by
16670b4f0baa set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents: 36190
diff changeset
   161
   counteracting the presence of "hAPP". *)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37480
diff changeset
   162
val spass_config : prover_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
   163
  {exec = ("ISABELLE_ATP", "scripts/spass"),
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
   164
   required_execs = [("SPASS_HOME", "SPASS")],
37514
b147d01b8ebc if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents: 37509
diff changeset
   165
   arguments = fn complete => fn timeout =>
37962
d7dbe01f48d7 keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents: 37926
diff changeset
   166
     ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
   167
      \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout))
37514
b147d01b8ebc if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents: 37509
diff changeset
   168
     |> not complete ? prefix "-SOS=1 ",
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
   169
   has_incomplete_mode = true,
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   170
   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
   171
   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
   172
     known_perl_failures @
37413
e856582fe9c4 improve ATP-specific error messages
blanchet
parents: 37347
diff changeset
   173
     [(IncompleteUnprovable, "SPASS beiseite: Completion found"),
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   174
      (TimedOut, "SPASS beiseite: Ran out of time"),
36965
67ae217c6b5c identify common SPASS error more clearly
blanchet
parents: 36924
diff changeset
   175
      (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
37413
e856582fe9c4 improve ATP-specific error messages
blanchet
parents: 37347
diff changeset
   176
      (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
   177
      (MalformedInput, "Free Variable"),
38519
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
   178
      (SpassTooOld, "tptp2dfg")],
38685
87a1e97a3ef3 adjust fudge factors in the light of experiments
blanchet
parents: 38680
diff changeset
   179
   default_max_relevant_per_iter = 40 (* FUDGE *),
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   180
   default_theory_relevant = true,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   181
   explicit_forall = true,
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   182
   use_conjecture_for_hypotheses = true}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   183
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   184
val spass = ("spass", spass_config)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   185
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   186
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   187
(* Vampire *)
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   188
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   189
val vampire_config : prover_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
   190
  {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
   191
   required_execs = [],
37514
b147d01b8ebc if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents: 37509
diff changeset
   192
   arguments = fn _ => fn timeout =>
38033
df99f022751d support latest version of Vampire (1.0) locally
blanchet
parents: 38032
diff changeset
   193
     "--mode casc -t " ^ string_of_int (to_generous_secs timeout) ^
38588
6a5b104f92cb thank Andrei instead of Tanya
blanchet
parents: 38519
diff changeset
   194
     " --thanks Andrei --input_file",
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
   195
   has_incomplete_mode = false,
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   196
   proof_delims =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   197
     [("=========== Refutation ==========",
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   198
       "======= End of refutation ======="),
38033
df99f022751d support latest version of Vampire (1.0) locally
blanchet
parents: 38032
diff changeset
   199
      ("% SZS output start Refutation", "% SZS output end Refutation"),
df99f022751d support latest version of Vampire (1.0) locally
blanchet
parents: 38032
diff changeset
   200
      ("% 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
   201
   known_failures =
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   202
     [(Unprovable, "UNPROVABLE"),
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   203
      (IncompleteUnprovable, "CANNOT PROVE"),
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
   204
      (TimedOut, "SZS status Timeout"),
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   205
      (Unprovable, "Satisfiability detected"),
38647
5500241da479 play with fudge factor + parse one more Vampire error
blanchet
parents: 38646
diff changeset
   206
      (Unprovable, "Termination reason: Satisfiable"),
38519
0dabf05fc86b rename enum values
blanchet
parents: 38495
diff changeset
   207
      (VampireTooOld, "not a valid option")],
38646
8fe717f5048e fiddle a bit with the SPASS fudge number
blanchet
parents: 38645
diff changeset
   208
   default_max_relevant_per_iter = 45 (* FUDGE *),
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   209
   default_theory_relevant = false,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   210
   explicit_forall = false,
38680
634a6d400c2e revert unintended change
blanchet
parents: 38678
diff changeset
   211
   use_conjecture_for_hypotheses = true}
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   212
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   213
val vampire = ("vampire", vampire_config)
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   214
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   215
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   216
(* Remote prover invocation via SystemOnTPTP *)
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   217
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   218
val systems = Synchronized.var "atp_systems" ([] : string list)
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   219
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   220
fun get_systems () =
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   221
  case bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   222
    (answer, 0) => split_lines answer
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
   223
  | (answer, _) =>
38065
9069e1ad1527 improved ATP error handling some more
blanchet
parents: 38064
diff changeset
   224
    error (case known_failure_in_output answer known_perl_failures of
9069e1ad1527 improved ATP error handling some more
blanchet
parents: 38064
diff changeset
   225
             SOME failure => string_for_failure failure
9069e1ad1527 improved ATP error handling some more
blanchet
parents: 38064
diff changeset
   226
           | NONE => perhaps (try (unsuffix "\n")) answer ^ ".")
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   227
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
   228
fun refresh_systems_on_tptp () =
37509
f39464d971c4 factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents: 37506
diff changeset
   229
  Synchronized.change systems (fn _ => get_systems ())
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   230
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
   231
fun find_system name [] systems = find_first (String.isPrefix name) 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
   232
  | 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
   233
    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
   234
      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
   235
    | 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
   236
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   237
fun get_system name versions =
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   238
  Synchronized.change_result systems
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   239
      (fn systems => (if null systems then get_systems () else 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
   240
                     |> `(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
   241
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
   242
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
   243
  case get_system name versions 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
   244
    NONE => error ("System " ^ quote name ^ " not available at SystemOnTPTP.")
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   245
  | SOME sys => sys
31835
b686d4df54c2 check for current versions on server
immler@in.tum.de
parents: 31832
diff changeset
   246
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
   247
fun remote_config system_name system_versions proof_delims known_failures
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   248
                  default_max_relevant_per_iter default_theory_relevant
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   249
                  use_conjecture_for_hypotheses =
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
   250
  {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
   251
   required_execs = [],
37514
b147d01b8ebc if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents: 37509
diff changeset
   252
   arguments = fn _ => fn timeout =>
38041
3b80d6082131 remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents: 38033
diff changeset
   253
     " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^
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
   254
     the_system system_name system_versions,
38645
4d5bbec1a598 be more generous towards SPASS's -SOS mode
blanchet
parents: 38631
diff changeset
   255
   has_incomplete_mode = false,
36369
d2cd0d04b8e6 handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents: 36289
diff changeset
   256
   proof_delims = insert (op =) tstp_proof_delims proof_delims,
38061
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   257
   known_failures =
685d1f0f75b3 handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents: 38049
diff changeset
   258
     known_failures @ known_perl_failures @
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38092
diff changeset
   259
     [(TimedOut, "says Timeout")],
38589
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   260
   default_max_relevant_per_iter = default_max_relevant_per_iter,
b03f8fe043ec added "max_relevant_per_iter" option to Sledgehammer
blanchet
parents: 38588
diff changeset
   261
   default_theory_relevant = default_theory_relevant,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   262
   explicit_forall = true,
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   263
   use_conjecture_for_hypotheses = use_conjecture_for_hypotheses}
28596
fcd463a6b6de tuned interfaces -- plain prover function, without thread;
wenzelm
parents: 28592
diff changeset
   264
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
   265
fun remotify_config system_name system_versions
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   266
        ({proof_delims, known_failures, default_max_relevant_per_iter,
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   267
          default_theory_relevant, use_conjecture_for_hypotheses, ...}
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   268
         : prover_config) : prover_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
   269
  remote_config system_name system_versions proof_delims known_failures
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   270
                default_max_relevant_per_iter default_theory_relevant
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   271
                use_conjecture_for_hypotheses
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   272
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   273
val remotify_name = prefix "remote_"
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
   274
fun remote_prover name system_name system_versions proof_delims known_failures
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   275
                  default_max_relevant_per_iter default_theory_relevant
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   276
                  use_conjecture_for_hypotheses =
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   277
  (remotify_name 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
   278
   remote_config system_name system_versions proof_delims known_failures
38631
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   279
                 default_max_relevant_per_iter default_theory_relevant
979a0b37f981 prefer TPTP "conjecture" tag to "hypothesis" on ATPs where this is possible;
blanchet
parents: 38603
diff changeset
   280
                 use_conjecture_for_hypotheses)
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
   281
fun remotify_prover (name, config) system_name system_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
   282
  (remotify_name name, remotify_config system_name system_versions config)
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
   283
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
   284
val remote_e = remotify_prover e "EP" ["1.0", "1.1", "1.2"]
38a926e033ad make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents: 38685
diff changeset
   285
val remote_vampire = remotify_prover vampire "Vampire" ["9.9", "0.6", "1.0"]
38603
a57d04dd1b25 fix SInE's error handling + run "vampire" locally if either SPASS or E is missing
blanchet
parents: 38598
diff changeset
   286
val remote_sine_e =
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
   287
  remote_prover "sine_e" "SInE" [] []
38646
8fe717f5048e fiddle a bit with the SPASS fudge number
blanchet
parents: 38645
diff changeset
   288
                [(Unprovable, "says Unknown")] 150 (* FUDGE *) false true
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   289
val remote_snark =
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
   290
  remote_prover "snark" "SNARK---" [] [("refutation.", "end_refutation.")] []
38646
8fe717f5048e fiddle a bit with the SPASS fudge number
blanchet
parents: 38645
diff changeset
   291
                50 (* FUDGE *) false true
38454
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   292
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   293
(* Setup *)
9043eefe8d71 detect old Vampire and give a nicer error message
blanchet
parents: 38433
diff changeset
   294
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
   295
fun is_installed ({exec, required_execs, ...} : prover_config) =
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
   296
  forall (curry (op <>) "" o getenv o fst) (exec :: required_execs)
38041
3b80d6082131 remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents: 38033
diff changeset
   297
fun maybe_remote (name, config) =
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   298
  name |> not (is_installed config) ? remotify_name
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   299
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   300
fun default_atps_param_value () =
38041
3b80d6082131 remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents: 38033
diff changeset
   301
  space_implode " " ([maybe_remote e] @
3b80d6082131 remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents: 38033
diff changeset
   302
                     (if is_installed (snd spass) then [fst spass] else []) @
38603
a57d04dd1b25 fix SInE's error handling + run "vampire" locally if either SPASS or E is missing
blanchet
parents: 38598
diff changeset
   303
                     [if forall (is_installed o snd) [e, spass] then
a57d04dd1b25 fix SInE's error handling + run "vampire" locally if either SPASS or E is missing
blanchet
parents: 38598
diff changeset
   304
                        remotify_name (fst vampire)
a57d04dd1b25 fix SInE's error handling + run "vampire" locally if either SPASS or E is missing
blanchet
parents: 38598
diff changeset
   305
                      else
a57d04dd1b25 fix SInE's error handling + run "vampire" locally if either SPASS or E is missing
blanchet
parents: 38598
diff changeset
   306
                        maybe_remote vampire,
a57d04dd1b25 fix SInE's error handling + run "vampire" locally if either SPASS or E is missing
blanchet
parents: 38598
diff changeset
   307
                      fst remote_sine_e])
36371
8c83ea1a7740 move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents: 36370
diff changeset
   308
38598
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   309
val provers = [e, spass, vampire, remote_e, remote_vampire, remote_sine_e,
ce117ef51999 added remote SInE and remote SNARK
blanchet
parents: 38596
diff changeset
   310
               remote_snark]
38023
962b0a7f544b more refactoring
blanchet
parents: 38022
diff changeset
   311
val setup = fold add_prover provers
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35865
diff changeset
   312
28592
824f8390aaa2 renamed AtpThread to AtpWrapper;
wenzelm
parents:
diff changeset
   313
end;