blanchet@38293: (* Title: HOL/Tools/ATP/atp_systems.ML wenzelm@28592: Author: Fabian Immler, TU Muenchen blanchet@36371: Author: Jasmin Blanchette, TU Muenchen wenzelm@28592: blanchet@36376: Setup for supported ATPs. wenzelm@28592: *) wenzelm@28592: blanchet@36376: signature ATP_SYSTEMS = wenzelm@28592: sig blanchet@43448: type format = ATP_Problem.format blanchet@43448: type formula_kind = ATP_Problem.formula_kind blanchet@39731: type failure = ATP_Proof.failure blanchet@38257: blanchet@40240: type atp_config = blanchet@43449: {exec : string * string, blanchet@43449: required_execs : (string * string) list, blanchet@43517: arguments : blanchet@44344: Proof.context -> bool -> string -> Time.time blanchet@44220: -> (unit -> (string * real) list) -> string, blanchet@43449: proof_delims : (string * string) list, blanchet@43449: known_failures : (failure * string) list, blanchet@43580: conj_sym_kind : formula_kind, blanchet@43580: prem_kind : formula_kind, blanchet@43449: formats : format list, blanchet@44344: best_slices : blanchet@44431: Proof.context -> (real * (bool * (int * string * string))) list} blanchet@38257: blanchet@44970: val force_sos : bool Config.T blanchet@44428: val e_smartN : string blanchet@44428: val e_autoN : string blanchet@44428: val e_fun_weightN : string blanchet@44428: val e_sym_offset_weightN : string blanchet@43517: val e_weight_method : string Config.T blanchet@43517: val e_default_fun_weight : real Config.T blanchet@43517: val e_fun_weight_base : real Config.T blanchet@43517: val e_fun_weight_span : real Config.T blanchet@43517: val e_default_sym_offs_weight : real Config.T blanchet@43517: val e_sym_offs_weight_base : real Config.T blanchet@43517: val e_sym_offs_weight_span : real Config.T blanchet@40240: val eN : string blanchet@40240: val spassN : string blanchet@40240: val vampireN : string blanchet@43803: val leo2N : string blanchet@43803: val satallaxN : string blanchet@44963: val e_sineN : string blanchet@40240: val snarkN : string blanchet@44963: val e_tofofN : string blanchet@43779: val waldmeisterN : string blanchet@42609: val z3_atpN : string blanchet@40241: val remote_prefix : string blanchet@42609: val remote_atp : blanchet@42609: string -> string -> string list -> (string * string) list blanchet@43580: -> (failure * string) list -> formula_kind -> formula_kind -> format list blanchet@44431: -> (Proof.context -> int * string) -> string * atp_config blanchet@40240: val add_atp : string * atp_config -> theory -> theory blanchet@40240: val get_atp : theory -> string -> atp_config blanchet@42591: val supported_atps : theory -> string list blanchet@40240: val is_atp_installed : theory -> string -> bool blanchet@35867: val refresh_systems_on_tptp : unit -> unit blanchet@35867: val setup : theory -> theory wenzelm@28592: end; wenzelm@28592: blanchet@36376: structure ATP_Systems : ATP_SYSTEMS = wenzelm@28592: struct wenzelm@28596: blanchet@43448: open ATP_Problem blanchet@39731: open ATP_Proof blanchet@39731: blanchet@40240: (* ATP configuration *) blanchet@35826: blanchet@40240: type atp_config = blanchet@43449: {exec : string * string, blanchet@43449: required_execs : (string * string) list, blanchet@43517: arguments : blanchet@44344: Proof.context -> bool -> string -> Time.time blanchet@44344: -> (unit -> (string * real) list) -> string, blanchet@43449: proof_delims : (string * string) list, blanchet@43449: known_failures : (failure * string) list, blanchet@43580: conj_sym_kind : formula_kind, blanchet@43580: prem_kind : formula_kind, blanchet@43449: formats : format list, blanchet@44344: best_slices : blanchet@44431: Proof.context -> (real * (bool * (int * string * string))) list} boehmes@32864: blanchet@43588: (* "best_slices" must be found empirically, taking a wholistic approach since blanchet@43588: the ATPs are run in parallel. The "real" components give the faction of the blanchet@44431: time available given to the slice and should add up to 1.0. The "bool" blanchet@43588: component indicates whether the slice's strategy is complete; the "int", the blanchet@44431: preferred number of facts to pass; the first "string", the preferred type blanchet@44437: system (which should be sound or quasi-sound); the second "string", extra blanchet@44437: information to the prover (e.g., SOS or no SOS). blanchet@43588: blanchet@43588: The last slice should be the most "normal" one, because it will get all the blanchet@44431: time available if the other slices fail early and also because it is used if blanchet@44431: slicing is disabled (e.g., by the minimizer). *) blanchet@43581: blanchet@38307: val known_perl_failures = blanchet@38340: [(CantConnect, "HTTP error"), blanchet@38340: (NoPerl, "env: perl"), blanchet@38311: (NoLibwwwPerl, "Can't locate HTTP")] wenzelm@28596: blanchet@40240: (* named ATPs *) blanchet@40240: blanchet@40240: val eN = "e" blanchet@44970: val leo2N = "leo2" blanchet@44970: val satallaxN = "satallax" blanchet@40240: val spassN = "spass" blanchet@40240: val vampireN = "vampire" blanchet@42611: val z3_atpN = "z3_atp" blanchet@44963: val e_sineN = "e_sine" blanchet@40240: val snarkN = "snark" blanchet@44963: val e_tofofN = "e_tofof" blanchet@43779: val waldmeisterN = "waldmeister" blanchet@40241: val remote_prefix = "remote_" wenzelm@28596: blanchet@38257: structure Data = Theory_Data blanchet@38257: ( blanchet@40240: type T = (atp_config * stamp) Symtab.table blanchet@38257: val empty = Symtab.empty blanchet@38257: val extend = I blanchet@38257: fun merge data : T = Symtab.merge (eq_snd op =) data blanchet@38257: handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") blanchet@38257: ) blanchet@36370: blanchet@44852: fun to_secs min time = Int.max (min, (Time.toMilliseconds time + 999) div 1000) blanchet@36140: blanchet@44344: val sosN = "sos" blanchet@44344: val no_sosN = "no_sos" blanchet@44344: blanchet@44970: val force_sos = Attrib.setup_config_bool @{binding atp_force_sos} (K false) blanchet@44970: blanchet@39731: blanchet@40240: (* E *) wenzelm@28596: blanchet@36369: val tstp_proof_delims = blanchet@43803: [("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation"), blanchet@43803: ("% SZS output start CNFRefutation", "% SZS output end CNFRefutation")] blanchet@36369: blanchet@44344: val e_smartN = "smart" blanchet@43517: val e_autoN = "auto" blanchet@43517: val e_fun_weightN = "fun_weight" blanchet@43517: val e_sym_offset_weightN = "sym_offset_weight" blanchet@42589: blanchet@43517: val e_weight_method = blanchet@44344: Attrib.setup_config_string @{binding atp_e_weight_method} (K e_smartN) blanchet@42641: (* FUDGE *) blanchet@43517: val e_default_fun_weight = blanchet@43517: Attrib.setup_config_real @{binding atp_e_default_fun_weight} (K 20.0) blanchet@43517: val e_fun_weight_base = blanchet@43517: Attrib.setup_config_real @{binding atp_e_fun_weight_base} (K 0.0) blanchet@43517: val e_fun_weight_span = blanchet@43517: Attrib.setup_config_real @{binding atp_e_fun_weight_span} (K 40.0) blanchet@43517: val e_default_sym_offs_weight = blanchet@43517: Attrib.setup_config_real @{binding atp_e_default_sym_offs_weight} (K 1.0) blanchet@43517: val e_sym_offs_weight_base = blanchet@43517: Attrib.setup_config_real @{binding atp_e_sym_offs_weight_base} (K ~20.0) blanchet@43517: val e_sym_offs_weight_span = blanchet@43517: Attrib.setup_config_real @{binding atp_e_sym_offs_weight_span} (K 60.0) blanchet@42589: blanchet@43314: fun e_weight_method_case method fw sow = blanchet@43517: if method = e_fun_weightN then fw blanchet@43517: else if method = e_sym_offset_weightN then sow blanchet@44349: else raise Fail ("unexpected " ^ quote method) blanchet@42589: blanchet@43517: fun scaled_e_weight ctxt method w = blanchet@43517: w * Config.get ctxt blanchet@43517: (e_weight_method_case method e_fun_weight_span e_sym_offs_weight_span) blanchet@43517: + Config.get ctxt blanchet@43517: (e_weight_method_case method e_fun_weight_base e_sym_offs_weight_base) blanchet@42589: |> Real.ceil |> signed_string_of_int blanchet@41561: blanchet@43517: fun e_weight_arguments ctxt method weights = blanchet@43517: if method = e_autoN then blanchet@42589: "-xAutoDev" blanchet@42589: else blanchet@44489: (* supplied by Stephan Schulz *) blanchet@41562: "--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \ blanchet@41562: \--destructive-er-aggressive --destructive-er --presat-simplify \ blanchet@41562: \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ blanchet@41562: \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \ blanchet@43314: \-H'(4*" ^ e_weight_method_case method "FunWeight" "SymOffsetWeight" ^ blanchet@42589: "(SimulateSOS, " ^ blanchet@43517: (e_weight_method_case method e_default_fun_weight e_default_sym_offs_weight blanchet@43517: |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^ blanchet@41562: ",20,1.5,1.5,1" ^ blanchet@43517: (weights () blanchet@43517: |> map (fn (s, w) => "," ^ s ^ ":" ^ scaled_e_weight ctxt method w) blanchet@43517: |> implode) ^ blanchet@41562: "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ blanchet@41562: \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\ blanchet@41562: \FIFOWeight(PreferProcessed))'" blanchet@41561: blanchet@43314: fun is_old_e_version () = (string_ord (getenv "E_VERSION", "1.2w") = LESS) blanchet@43314: blanchet@43517: fun effective_e_weight_method ctxt = blanchet@43517: if is_old_e_version () then e_autoN else Config.get ctxt e_weight_method blanchet@43314: blanchet@40240: val e_config : atp_config = blanchet@38338: {exec = ("E_HOME", "eproof"), blanchet@38338: required_execs = [], blanchet@44220: arguments = blanchet@44429: fn ctxt => fn _ => fn method => fn timeout => fn weights => blanchet@44344: "--tstp-in --tstp-out -l5 " ^ e_weight_arguments ctxt method weights ^ blanchet@44852: " -tAutoDev --silent --cpu-limit=" ^ string_of_int (to_secs 2 timeout), blanchet@43803: proof_delims = tstp_proof_delims, blanchet@36265: known_failures = blanchet@38229: [(Unprovable, "SZS status: CounterSatisfiable"), blanchet@38229: (Unprovable, "SZS status CounterSatisfiable"), blanchet@43713: (ProofMissing, "SZS status Theorem"), blanchet@36370: (TimedOut, "Failure: Resource limit exceeded (time)"), blanchet@36370: (TimedOut, "time limit exceeded"), blanchet@44328: (OutOfResources, "# Cannot determine problem status"), blanchet@36370: (OutOfResources, "SZS status: ResourceOut"), blanchet@36370: (OutOfResources, "SZS status ResourceOut")], blanchet@44327: conj_sym_kind = Hypothesis, blanchet@43580: prem_kind = Conjecture, blanchet@43778: formats = [FOF], blanchet@43517: best_slices = fn ctxt => blanchet@44344: let val method = effective_e_weight_method ctxt in blanchet@44345: (* FUDGE *) blanchet@44344: if method = e_smartN then blanchet@44431: [(0.333, (true, (500, "mangled_tags?", e_fun_weightN))), blanchet@44860: (0.334, (true, (50, "mangled_guards?", e_fun_weightN))), blanchet@44431: (0.333, (true, (1000, "mangled_tags?", e_sym_offset_weightN)))] blanchet@44344: else blanchet@44431: [(1.0, (true, (500, "mangled_tags?", method)))] blanchet@44344: end} blanchet@38698: blanchet@40240: val e = (eN, e_config) wenzelm@28596: wenzelm@28596: blanchet@44970: (* LEO-II *) blanchet@44970: blanchet@44970: val leo2_config : atp_config = blanchet@44970: {exec = ("LEO2_HOME", "leo"), blanchet@44970: required_execs = [], blanchet@44970: arguments = blanchet@44970: fn _ => fn _ => fn sos => fn timeout => fn _ => blanchet@44970: "--proofoutput --timeout " ^ string_of_int (to_secs 1 timeout) blanchet@44970: |> sos = sosN ? prefix "--sos ", blanchet@44970: proof_delims = tstp_proof_delims, blanchet@44970: known_failures = [], blanchet@44970: conj_sym_kind = Axiom, blanchet@44970: prem_kind = Hypothesis, blanchet@44970: formats = [THF, FOF], blanchet@44970: best_slices = fn ctxt => blanchet@44970: (* FUDGE *) blanchet@44970: [(0.667, (false, (150, "simple_higher", sosN))), blanchet@44970: (0.333, (true, (50, "simple_higher", no_sosN)))] blanchet@44970: |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single blanchet@44970: else I)} blanchet@44970: blanchet@44970: val leo2 = (leo2N, leo2_config) blanchet@44970: blanchet@44970: blanchet@44970: (* Satallax *) blanchet@44970: blanchet@44970: val satallax_config : atp_config = blanchet@44970: {exec = ("SATALLAX_HOME", "satallax"), blanchet@44970: required_execs = [], blanchet@44970: arguments = blanchet@44970: fn _ => fn _ => fn _ => fn timeout => fn _ => blanchet@44970: "-t " ^ string_of_int (to_secs 1 timeout), blanchet@44970: proof_delims = tstp_proof_delims, blanchet@44970: known_failures = [(ProofMissing, "SZS status Theorem")], blanchet@44970: conj_sym_kind = Axiom, blanchet@44970: prem_kind = Hypothesis, blanchet@44970: formats = [THF], blanchet@44970: best_slices = K [(1.0, (true, (100, "simple_higher", "")))] (* FUDGE *)} blanchet@44970: blanchet@44970: val satallax = (satallaxN, satallax_config) blanchet@44970: blanchet@44970: blanchet@39731: (* SPASS *) blanchet@39731: blanchet@36219: (* The "-VarWeight=3" option helps the higher-order problems, probably by blanchet@36219: counteracting the presence of "hAPP". *) blanchet@40240: val spass_config : atp_config = blanchet@38338: {exec = ("ISABELLE_ATP", "scripts/spass"), blanchet@39246: required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")], blanchet@44431: arguments = fn _ => fn _ => fn sos => fn timeout => fn _ => blanchet@38199: ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ blanchet@44852: \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 1 timeout)) blanchet@44344: |> sos = sosN ? prefix "-SOS=1 ", blanchet@36369: proof_delims = [("Here is a proof", "Formulae used in the proof")], blanchet@36289: known_failures = blanchet@38307: known_perl_failures @ blanchet@43891: [(GaveUp, "SPASS beiseite: Completion found"), blanchet@36370: (TimedOut, "SPASS beiseite: Ran out of time"), blanchet@36958: (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), blanchet@37388: (MalformedInput, "Undefined symbol"), blanchet@37389: (MalformedInput, "Free Variable"), blanchet@39490: (SpassTooOld, "tptp2dfg"), blanchet@39490: (InternalError, "Please report this error")], blanchet@44327: conj_sym_kind = Hypothesis, blanchet@43580: prem_kind = Conjecture, blanchet@43778: formats = [FOF], blanchet@43590: best_slices = fn ctxt => blanchet@43588: (* FUDGE *) blanchet@44431: [(0.333, (false, (150, "mangled_tags", sosN))), blanchet@44431: (0.333, (false, (300, "poly_tags?", sosN))), blanchet@44431: (0.334, (true, (50, "mangled_tags?", no_sosN)))] blanchet@44970: |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single blanchet@43590: else I)} blanchet@38698: blanchet@40240: val spass = (spassN, spass_config) wenzelm@28596: blanchet@38698: blanchet@37509: (* Vampire *) blanchet@37509: blanchet@40240: val vampire_config : atp_config = blanchet@38338: {exec = ("VAMPIRE_HOME", "vampire"), blanchet@38338: required_execs = [], blanchet@44431: arguments = fn _ => fn _ => fn sos => fn timeout => fn _ => blanchet@44852: "--proof tptp --mode casc -t " ^ string_of_int (to_secs 1 timeout) ^ blanchet@41451: " --thanks \"Andrei and Krystof\" --input_file" blanchet@44344: |> sos = sosN ? prefix "--sos on ", blanchet@37509: proof_delims = blanchet@37509: [("=========== Refutation ==========", blanchet@37509: "======= End of refutation ======="), blanchet@38279: ("% SZS output start Refutation", "% SZS output end Refutation"), blanchet@38279: ("% SZS output start Proof", "% SZS output end Proof")], blanchet@37509: known_failures = blanchet@43891: [(GaveUp, "UNPROVABLE"), blanchet@43891: (GaveUp, "CANNOT PROVE"), blanchet@43891: (GaveUp, "SZS status GaveUp"), blanchet@43751: (ProofIncomplete, "predicate_definition_introduction,[]"), blanchet@38338: (TimedOut, "SZS status Timeout"), blanchet@37509: (Unprovable, "Satisfiability detected"), blanchet@38885: (Unprovable, "Termination reason: Satisfiable"), blanchet@39490: (VampireTooOld, "not a valid option"), blanchet@39490: (Interrupted, "Aborted by signal SIGINT")], blanchet@44327: conj_sym_kind = Conjecture, blanchet@43580: prem_kind = Conjecture, blanchet@43778: formats = [FOF], blanchet@43590: best_slices = fn ctxt => blanchet@43588: (* FUDGE *) blanchet@44860: [(0.333, (false, (150, "poly_guards", sosN))), blanchet@44860: (0.334, (true, (50, "mangled_guards?", no_sosN))), blanchet@44431: (0.333, (false, (500, "mangled_tags?", sosN)))] blanchet@44970: |> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single blanchet@43590: else I)} blanchet@38698: blanchet@40240: val vampire = (vampireN, vampire_config) blanchet@37509: blanchet@38698: blanchet@42611: (* Z3 with TPTP syntax *) blanchet@42611: blanchet@42611: val z3_atp_config : atp_config = blanchet@42611: {exec = ("Z3_HOME", "z3"), blanchet@42611: required_execs = [], blanchet@44220: arguments = fn _ => fn _ => fn _ => fn timeout => fn _ => blanchet@44852: "MBQI=true -p -t:" ^ string_of_int (to_secs 1 timeout), blanchet@42611: proof_delims = [], blanchet@42611: known_failures = blanchet@42613: [(Unprovable, "\nsat"), blanchet@43891: (GaveUp, "\nunknown"), blanchet@43891: (GaveUp, "SZS status Satisfiable"), blanchet@43314: (ProofMissing, "\nunsat"), blanchet@43314: (ProofMissing, "SZS status Unsatisfiable")], blanchet@43580: conj_sym_kind = Hypothesis, blanchet@43580: prem_kind = Hypothesis, blanchet@43778: formats = [FOF], blanchet@43588: best_slices = blanchet@43588: (* FUDGE (FIXME) *) blanchet@44860: K [(0.5, (false, (250, "mangled_guards?", ""))), blanchet@44860: (0.25, (false, (125, "mangled_guards?", ""))), blanchet@44860: (0.125, (false, (62, "mangled_guards?", ""))), blanchet@44860: (0.125, (false, (31, "mangled_guards?", "")))]} blanchet@42611: blanchet@42611: val z3_atp = (z3_atpN, z3_atp_config) blanchet@42611: blanchet@42611: blanchet@40240: (* Remote ATP invocation via SystemOnTPTP *) wenzelm@28596: blanchet@38307: val systems = Synchronized.var "atp_systems" ([] : string list) immler@31828: immler@31828: fun get_systems () = wenzelm@44721: case Isabelle_System.bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of blanchet@39731: (output, 0) => split_lines output blanchet@39731: | (output, _) => blanchet@39731: error (case extract_known_failure known_perl_failures output of blanchet@42615: SOME failure => string_for_failure failure blanchet@39731: | NONE => perhaps (try (unsuffix "\n")) output ^ ".") immler@31828: blanchet@43408: fun find_system name [] systems = blanchet@43408: find_first (String.isPrefix (name ^ "---")) systems blanchet@38929: | find_system name (version :: versions) systems = blanchet@38929: case find_first (String.isPrefix (name ^ "---" ^ version)) systems of blanchet@38929: NONE => find_system name versions systems blanchet@38929: | res => res blanchet@38929: blanchet@38929: fun get_system name versions = blanchet@38812: Synchronized.change_result systems blanchet@38812: (fn systems => (if null systems then get_systems () else systems) blanchet@43796: |> `(`(find_system name versions))) immler@31828: blanchet@38929: fun the_system name versions = blanchet@38929: case get_system name versions of blanchet@43796: (SOME sys, _) => sys blanchet@43796: | (NONE, []) => error ("SystemOnTPTP is currently not available.") blanchet@43796: | (NONE, syss) => blanchet@43796: error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n" ^ blanchet@43796: "(Available systems: " ^ commas_quote syss ^ ".)") wenzelm@28596: blanchet@41396: val max_remote_secs = 240 (* give Geoff Sutcliffe's servers a break *) blanchet@41396: blanchet@38929: fun remote_config system_name system_versions proof_delims known_failures blanchet@43588: conj_sym_kind prem_kind formats best_slice : atp_config = blanchet@38338: {exec = ("ISABELLE_ATP", "scripts/remote_atp"), blanchet@38338: required_execs = [], blanchet@44220: arguments = fn _ => fn _ => fn _ => fn timeout => fn _ => blanchet@44852: "-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout)) blanchet@41396: ^ " -s " ^ the_system system_name system_versions, blanchet@43803: proof_delims = union (op =) tstp_proof_delims proof_delims, blanchet@43806: known_failures = known_failures @ known_perl_failures @ blanchet@43782: [(Unprovable, "says Satisfiable"), blanchet@43840: (Unprovable, "says CounterSatisfiable"), blanchet@43891: (GaveUp, "says Unknown"), blanchet@43891: (GaveUp, "says GaveUp"), blanchet@43806: (ProofMissing, "says Theorem"), blanchet@43840: (ProofMissing, "says Unsatisfiable"), blanchet@43794: (TimedOut, "says Timeout"), blanchet@43794: (Inappropriate, "says Inappropriate")], blanchet@43580: conj_sym_kind = conj_sym_kind, blanchet@43580: prem_kind = prem_kind, blanchet@43449: formats = formats, blanchet@44344: best_slices = fn ctxt => blanchet@44493: let val (max_relevant, type_enc) = best_slice ctxt in blanchet@44493: [(1.0, (false, (max_relevant, type_enc, "")))] blanchet@44344: end} blanchet@43314: blanchet@44371: fun remotify_config system_name system_versions best_slice blanchet@44371: ({proof_delims, known_failures, conj_sym_kind, prem_kind, formats, ...} blanchet@44371: : atp_config) : atp_config = blanchet@38929: remote_config system_name system_versions proof_delims known_failures blanchet@44371: conj_sym_kind prem_kind formats best_slice blanchet@38257: blanchet@40240: fun remote_atp name system_name system_versions proof_delims known_failures blanchet@43588: conj_sym_kind prem_kind formats best_slice = blanchet@40241: (remote_prefix ^ name, blanchet@38929: remote_config system_name system_versions proof_delims known_failures blanchet@43588: conj_sym_kind prem_kind formats best_slice) blanchet@44371: fun remotify_atp (name, config) system_name system_versions best_slice = blanchet@44371: (remote_prefix ^ name, blanchet@44371: remotify_config system_name system_versions best_slice config) wenzelm@28592: blanchet@44371: val remote_e = blanchet@44371: remotify_atp e "EP" ["1.0", "1.1", "1.2"] blanchet@44431: (K (750, "mangled_tags?") (* FUDGE *)) blanchet@44970: val remote_leo2 = blanchet@44970: remotify_atp leo2 "LEO-II" ["1.2.8", "1.2.6"] blanchet@44970: (K (100, "simple_higher") (* FUDGE *)) blanchet@44970: val remote_satallax = blanchet@44970: remotify_atp satallax "Satallax" ["2.1", "2.0", "2"] blanchet@44970: (K (100, "simple_higher") (* FUDGE *)) blanchet@44371: val remote_vampire = blanchet@44371: remotify_atp vampire "Vampire" ["0.6", "9.0", "1.0"] blanchet@44860: (K (200, "mangled_guards?") (* FUDGE *)) blanchet@44371: val remote_z3_atp = blanchet@44860: remotify_atp z3_atp "Z3" ["2.18"] (K (250, "mangled_guards?") (* FUDGE *)) blanchet@44963: val remote_e_sine = blanchet@44963: remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Axiom blanchet@44860: Conjecture [FOF] (K (500, "mangled_guards?") (* FUDGE *)) blanchet@42611: val remote_snark = blanchet@43780: remote_atp snarkN "SNARK" ["20080805r029", "20080805r024"] blanchet@43803: [("refutation.", "end_refutation.")] [] Hypothesis Hypothesis blanchet@44431: [TFF, FOF] (K (100, "simple") (* FUDGE *)) blanchet@44963: val remote_e_tofof = blanchet@44963: remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config) blanchet@44431: Axiom Hypothesis [TFF] (K (150, "simple") (* FUDGE *)) blanchet@43779: val remote_waldmeister = blanchet@43779: remote_atp waldmeisterN "Waldmeister" ["710"] blanchet@43785: [("#START OF PROOF", "Proved Goals:")] blanchet@43794: [(OutOfResources, "Too many function symbols"), blanchet@43794: (Crashed, "Unrecoverable Segmentation Fault")] blanchet@43794: Hypothesis Hypothesis [CNF_UEQ] blanchet@44431: (K (50, "mangled_tags?") (* FUDGE *)) blanchet@38698: blanchet@38698: (* Setup *) blanchet@38698: blanchet@40240: fun add_atp (name, config) thy = blanchet@40240: Data.map (Symtab.update_new (name, (config, stamp ()))) thy blanchet@40240: handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") blanchet@36371: blanchet@40240: fun get_atp thy name = blanchet@40240: the (Symtab.lookup (Data.get thy) name) |> fst blanchet@40240: handle Option.Option => error ("Unknown ATP: " ^ name ^ ".") blanchet@36371: blanchet@42591: val supported_atps = Symtab.keys o Data.get blanchet@40240: blanchet@40240: fun is_atp_installed thy name = blanchet@40240: let val {exec, required_execs, ...} = get_atp thy name in blanchet@40240: forall (curry (op <>) "" o getenv o fst) (exec :: required_execs) blanchet@40240: end blanchet@40240: blanchet@40240: fun refresh_systems_on_tptp () = blanchet@40240: Synchronized.change systems (fn _ => get_systems ()) blanchet@40240: blanchet@43803: val atps = blanchet@44970: [e, leo2, satallax, spass, vampire, z3_atp, remote_e, remote_leo2, blanchet@44970: remote_satallax, remote_vampire, remote_z3_atp, remote_e_sine, remote_snark, blanchet@44970: remote_e_tofof, remote_waldmeister] blanchet@40240: val setup = fold add_atp atps blanchet@35867: wenzelm@28592: end;