diff --git a/src/HOL/Tools/ATP/atp_proof.ML b/src/HOL/Tools/ATP/atp_proof.ML --- a/src/HOL/Tools/ATP/atp_proof.ML +++ b/src/HOL/Tools/ATP/atp_proof.ML @@ -1,730 +1,734 @@ (* Title: HOL/Tools/ATP/atp_proof.ML Author: Lawrence C. Paulson, Cambridge University Computer Laboratory Author: Claire Quigley, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Author: Mathias Fleury, ENS Rennes Abstract representation of ATP proofs and TSTP/SPASS syntax. *) signature ATP_PROOF = sig type 'a atp_type = 'a ATP_Problem.atp_type type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term type atp_formula_role = ATP_Problem.atp_formula_role type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula type 'a atp_problem = 'a ATP_Problem.atp_problem exception UNRECOGNIZED_ATP_PROOF of unit datatype atp_failure = MaybeUnprovable | Unprovable | GaveUp | ProofMissing | ProofIncomplete | ProofUnparsable | UnsoundProof of bool * string list | TimedOut | Inappropriate | OutOfResources | MalformedInput | MalformedOutput | Interrupted | Crashed | InternalError | UnknownError of string type atp_step_name = string * string list type ('a, 'b) atp_step = atp_step_name * atp_formula_role * 'a * 'b * atp_step_name list type 'a atp_proof = (('a, 'a, ('a, 'a atp_type) atp_term, 'a) atp_formula, string) atp_step list val agsyholN : string val alt_ergoN : string val cvc4N : string val eN : string val iproverN : string val leo2N : string val leo3N : string val satallaxN : string val spassN : string val vampireN : string val veritN : string val waldmeisterN : string val z3N : string val zipperpositionN : string val remote_prefix : string val dummy_fofN : string val dummy_tfxN : string val dummy_thfN : string val agsyhol_core_rule : string val spass_input_rule : string val spass_pre_skolemize_rule : string val spass_skolemize_rule : string val short_output : bool -> string -> string val string_of_atp_failure : atp_failure -> string val extract_important_message : string -> string val extract_known_atp_failure : (atp_failure * string) list -> string -> atp_failure option val extract_tstplike_proof_and_outcome : bool -> (string * string) list -> (atp_failure * string) list -> string -> string * atp_failure option val is_same_atp_step : atp_step_name -> atp_step_name -> bool val scan_general_id : string list -> string * string list val parse_fol_formula : string list -> (string, string atp_type, (string, string atp_type) atp_term, string) atp_formula * string list val clean_up_atp_proof_dependencies : string atp_proof -> string atp_proof val map_term_names_in_atp_proof : (string -> string) -> string atp_proof -> string atp_proof val nasty_atp_proof : string Symtab.table -> string atp_proof -> string atp_proof val skip_term : string list -> string * string list val parse_hol_formula : string list -> ('a, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, 'c) ATP_Problem.atp_formula * string list val dummy_atype : string ATP_Problem.atp_type val role_of_tptp_string : string -> ATP_Problem.atp_formula_role - val parse_line : string -> ('a * string ATP_Problem.atp_problem_line list) list -> + val parse_line : bool -> string -> ('a * string ATP_Problem.atp_problem_line list) list -> string list -> ((string * string list) * ATP_Problem.atp_formula_role * (string, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, 'c) ATP_Problem.atp_formula * string * (string * 'd list) list) list * string list val core_inference : 'a -> 'b -> ('b * 'b list) * ATP_Problem.atp_formula_role * ('c, 'd, (string, 'e) ATP_Problem.atp_term, 'f) ATP_Problem.atp_formula * 'a * 'g list val vampire_step_name_ord : (string * 'a) ord val core_of_agsyhol_proof : string -> string list option val string_of_atp_step : ('a -> string) -> ('b -> string) -> ('a, 'b) atp_step -> string - val atp_proof_of_tstplike_proof : string -> string atp_problem -> string -> string atp_proof + val atp_proof_of_tstplike_proof : bool -> string -> string atp_problem -> string -> + string atp_proof end; structure ATP_Proof : ATP_PROOF = struct open ATP_Util open ATP_Problem val agsyholN = "agsyhol" val alt_ergoN = "alt_ergo" val cvc4N = "cvc4" val eN = "e" val iproverN = "iprover" val leo2N = "leo2" val leo3N = "leo3" val satallaxN = "satallax" val spassN = "spass" val vampireN = "vampire" val veritN = "verit" val waldmeisterN = "waldmeister" val z3N = "z3" val zipperpositionN = "zipperposition" val remote_prefix = "remote_" val dummy_fofN = "dummy_fof" val dummy_tfxN = "dummy_tfx" val dummy_thfN = "dummy_thf" val agsyhol_core_rule = "__agsyhol_core" (* arbitrary *) val spass_input_rule = "Inp" val spass_pre_skolemize_rule = "__Sko0" (* arbitrary *) val spass_skolemize_rule = "__Sko" (* arbitrary *) exception UNRECOGNIZED_ATP_PROOF of unit datatype atp_failure = MaybeUnprovable | Unprovable | GaveUp | ProofMissing | ProofIncomplete | ProofUnparsable | UnsoundProof of bool * string list | TimedOut | Inappropriate | OutOfResources | MalformedInput | MalformedOutput | Interrupted | Crashed | InternalError | UnknownError of string fun short_output verbose output = if verbose then if output = "" then "No details available" else elide_string 1000 output else "" fun from_lemmas [] = "" | from_lemmas ss = " from " ^ space_implode " " (Try.serial_commas "and" (map quote ss)) fun string_of_atp_failure MaybeUnprovable = "Problem maybe unprovable" | string_of_atp_failure Unprovable = "Unprovable problem" | string_of_atp_failure GaveUp = "Gave up" | string_of_atp_failure ProofMissing = "Claims the conjecture is a theorem but did not provide a proof" | string_of_atp_failure ProofIncomplete = "Claims the conjecture is a theorem but provided an incomplete proof" | string_of_atp_failure ProofUnparsable = "Claims the conjecture is a theorem but provided an unparsable proof" | string_of_atp_failure (UnsoundProof (false, ss)) = "Derived the lemma \"False\"" ^ from_lemmas ss ^ ", probably due to the use of an unsound type encoding" | string_of_atp_failure (UnsoundProof (true, ss)) = "Derived the lemma \"False\"" ^ from_lemmas ss ^ ", which could be due to a bug in Sledgehammer or to inconsistent axioms (including \"sorry\"s)" | string_of_atp_failure TimedOut = "Timed out" | string_of_atp_failure Inappropriate = "Problem outside the prover's scope" | string_of_atp_failure OutOfResources = "Ran out of resources" | string_of_atp_failure MalformedInput = "Malformed problem" | string_of_atp_failure MalformedOutput = "Malformed output" | string_of_atp_failure Interrupted = "Interrupted" | string_of_atp_failure Crashed = "Crashed" | string_of_atp_failure InternalError = "Internal prover error" | string_of_atp_failure (UnknownError s) = "Prover error" ^ (if s = "" then " (pass the \"verbose\" option for details)" else ":\n" ^ s) fun extract_delimited (begin_delim, end_delim) output = (case first_field begin_delim output of SOME (_, tail) => (case first_field "\n" tail of SOME (_, tail') => if end_delim = "" then tail' else (case first_field end_delim tail' of SOME (body, _) => body | NONE => "") | NONE => "") | NONE => "") val tstp_important_message_delims = ("% SZS start RequiredInformation", "% SZS end RequiredInformation") fun extract_important_message output = (case extract_delimited tstp_important_message_delims output of "" => "" | s => s |> space_explode "\n" |> filter_out (curry (op =) "") |> map (perhaps (try (unprefix "%"))) |> map (perhaps (try (unprefix " "))) |> space_implode "\n " |> quote) (* Splits by the first possible of a list of delimiters. *) fun extract_tstplike_proof delims output = (case apply2 (find_first (fn s => String.isSubstring s output)) (ListPair.unzip delims) of (SOME begin_delim, SOME end_delim) => extract_delimited (begin_delim, end_delim) output | _ => "") fun extract_known_atp_failure known_failures output = known_failures |> find_first (fn (_, pattern) => String.isSubstring pattern output) |> Option.map fst fun extract_tstplike_proof_and_outcome verbose proof_delims known_failures output = let val known_atp_failure = extract_known_atp_failure known_failures output val tstplike_proof = extract_tstplike_proof proof_delims output in (case (tstplike_proof, known_atp_failure) of (_, SOME ProofIncomplete) => ("", NONE) | (_, SOME ProofUnparsable) => ("", NONE) | ("", SOME ProofMissing) => ("", NONE) | ("", NONE) => ("", SOME (UnknownError (short_output verbose output))) | res as ("", _) => res | (tstplike_proof, _) => (tstplike_proof, NONE)) end type atp_step_name = string * string list fun is_same_atp_step (s1, _) (s2, _) = s1 = s2 val vampire_fact_prefix = "f" fun vampire_step_name_ord p = let val q = apply2 fst p in (* The "unprefix" part is to cope with Vampire's output, which puts facts with names of the form "fN" where N is an integer in reverse order. *) (case apply2 (Int.fromString o perhaps (try (unprefix vampire_fact_prefix))) q of (SOME i, SOME j) => int_ord (i, j) | (SOME _, NONE) => LESS | (NONE, SOME _) => GREATER | (NONE, NONE) => string_ord q) end type ('a, 'b) atp_step = atp_step_name * atp_formula_role * 'a * 'b * atp_step_name list type 'a atp_proof = (('a, 'a, ('a, 'a atp_type) atp_term, 'a) atp_formula, string) atp_step list (**** PARSING OF TSTP FORMAT ****) (* Strings enclosed in single quotes (e.g., file names), identifiers possibly starting with "$" and possibly with "!" in them. *) val scan_general_id = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode || (Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig >> (op ^ o apply2 implode)) -- Scan.optional (Scan.repeat ($$ "!") -- Scan.many1 Symbol.is_letdig >> (op ^ o apply2 implode)) "" >> op ^ fun skip_term x = let fun skip _ accum [] = (accum, []) | skip n accum (ss as s :: ss') = if (s = "," orelse s = ".") andalso n = 0 then (accum, ss) else if member (op =) [")", "]"] s then if n = 0 then (accum, ss) else skip (n - 1) (s :: accum) ss' else if member (op =) ["(", "["] s then skip (n + 1) (s :: accum) ss' else skip n (s :: accum) ss' in (skip 0 [] #>> (rev #> implode)) x end and skip_terms x = (skip_term ::: Scan.repeat ($$ "," |-- skip_term)) x datatype source = File_Source of string * string option | Inference_Source of string * string list | Introduced_Source of string val dummy_phi = AAtom (ATerm (("", []), [])) val dummy_atype = AType (("", []), []) (* "skip_term" is there to cope with Waldmeister nonsense such as "theory(equality)". *) fun parse_dependency x = (parse_inference_source >> snd || scan_general_id --| skip_term >> single) x and parse_dependencies x = (Scan.repeats (Scan.option ($$ ",") |-- parse_dependency) >> (filter_out (curry (op =) "theory"))) x and parse_file_source x = (Scan.this_string "file" |-- $$ "(" |-- scan_general_id -- Scan.option ($$ "," |-- scan_general_id --| Scan.option ($$ "," |-- $$ "[" -- Scan.option scan_general_id --| $$ "]")) --| $$ ")") x and parse_inference_source x = (Scan.this_string "inference" |-- $$ "(" |-- scan_general_id --| skip_term --| $$ "," --| skip_term --| $$ "," --| $$ "[" -- parse_dependencies --| $$ "]" --| $$ ")") x and parse_introduced_source x = (Scan.this_string "introduced" |-- $$ "(" |-- scan_general_id --| Scan.option ($$ "," |-- skip_term) --| $$ ")") x and parse_source x = (parse_file_source >> File_Source >> SOME || parse_inference_source >> Inference_Source >> SOME || parse_introduced_source >> Introduced_Source >> SOME || scan_general_id >> (fn s => SOME (Inference_Source ("", [s]))) (* for E *) || skip_term >> K NONE) x fun list_app (f, args) = fold (fn arg => fn f => ATerm ((tptp_app, []), [f, arg])) args f fun parse_class x = scan_general_id x and parse_classes x = (parse_class ::: Scan.repeat ($$ "&" |-- parse_class)) x fun parse_type x = (($$ "(" |-- parse_type --| $$ ")" || Scan.this_string tptp_pi_binder |-- $$ "[" |-- skip_terms --| $$ "]" --| $$ ":" -- parse_type >> (fn (_, ty) => ty (* currently ignoring type constructor declarations anyway *)) || (scan_general_id -- Scan.optional ($$ "{" |-- parse_classes --| $$ "}") []) -- Scan.optional ($$ "(" |-- parse_types --| $$ ")") [] >> AType) -- Scan.option (($$ tptp_app || $$ tptp_fun_type || $$ tptp_product_type) -- parse_type) >> (fn (a, NONE) => a | (a, SOME (bin_op, b)) => if bin_op = tptp_app then (case a of AType (s_clss, tys) => AType (s_clss, tys @ [b]) | _ => raise UNRECOGNIZED_ATP_PROOF ()) else if bin_op = tptp_fun_type then AFun (a, b) else if bin_op = tptp_product_type then AType ((tptp_product_type, []), [a, b]) else raise Fail "impossible case")) x and parse_types x = (parse_type ::: Scan.repeat ($$ "," |-- parse_type)) x (* We currently half ignore types. *) fun parse_fol_optional_type_signature x = (Scan.option ($$ tptp_has_type |-- parse_type) >> (fn some as SOME (AType ((s, []), [])) => if s = dfg_individual_type then NONE else some | res => res)) x and parse_fol_arg x = ($$ "(" |-- parse_fol_term --| $$ ")" --| parse_fol_optional_type_signature || scan_general_id -- parse_fol_optional_type_signature -- Scan.optional ($$ "<" |-- parse_types --| $$ ">") [] -- Scan.optional ($$ "(" |-- parse_fol_terms --| $$ ")") [] >> (fn (((s, ty_opt), tyargs), args) => if is_tptp_variable s andalso null tyargs andalso null args andalso is_some ty_opt then ATerm ((s, the_list ty_opt), []) else ATerm ((s, tyargs), args))) x and parse_fol_term x = (parse_fol_arg -- Scan.repeat ($$ tptp_app |-- parse_fol_arg) --| parse_fol_optional_type_signature >> list_app) x and parse_fol_terms x = (parse_fol_term ::: Scan.repeat ($$ "," |-- parse_fol_term)) x fun parse_fol_atom x = (parse_fol_term -- Scan.option (Scan.option ($$ tptp_not_infix) --| $$ tptp_equal -- parse_fol_term) >> (fn (u1, NONE) => AAtom u1 | (u1, SOME (neg, u2)) => AAtom (ATerm (("equal", []), [u1, u2])) |> is_some neg ? mk_anot)) x (* TPTP formulas are fully parenthesized, so we don't need to worry about operator precedence. *) fun parse_fol_literal x = ((Scan.repeat ($$ tptp_not) >> length) -- ($$ "(" |-- parse_fol_formula --| $$ ")" || parse_fol_quantified_formula || parse_fol_atom) >> (fn (n, phi) => phi |> n mod 2 = 1 ? mk_anot)) x and parse_fol_formula x = (parse_fol_literal -- Scan.option ((Scan.this_string tptp_implies || Scan.this_string tptp_iff || Scan.this_string tptp_not_iff || Scan.this_string tptp_if || $$ tptp_or || $$ tptp_and) -- parse_fol_formula) >> (fn (phi1, NONE) => phi1 | (phi1, SOME (c, phi2)) => if c = tptp_implies then mk_aconn AImplies phi1 phi2 else if c = tptp_iff then mk_aconn AIff phi1 phi2 else if c = tptp_not_iff then mk_anot (mk_aconn AIff phi1 phi2) else if c = tptp_if then mk_aconn AImplies phi2 phi1 else if c = tptp_or then mk_aconn AOr phi1 phi2 else if c = tptp_and then mk_aconn AAnd phi1 phi2 else raise Fail ("impossible connective " ^ quote c))) x and parse_fol_quantified_formula x = (($$ tptp_forall >> K AForall || $$ tptp_exists >> K AExists) --| $$ "[" -- parse_fol_terms --| $$ "]" --| $$ ":" -- parse_fol_literal >> (fn ((q, ts), phi) => AQuant (q, map (fn ATerm ((s, _), _) => (s, NONE)) ts, phi))) x val parse_tstp_extra_arguments = Scan.optional ($$ "," |-- parse_source --| Scan.option ($$ "," |-- skip_term)) NONE val waldmeister_conjecture_name = "conjecture_1" fun is_same_term subst tm1 tm2 = let fun do_term_pair (AAbs (((var1, typ1), body1), args1)) (AAbs (((var2, typ2), body2), args2)) (SOME subst) = if typ1 <> typ2 andalso length args1 = length args2 then NONE else let val ls = length subst in SOME ((var1, var2) :: subst) |> do_term_pair body1 body2 |> (fn SOME subst => SOME (nth_drop (length subst - ls - 1) subst) | NONE => NONE) |> (if length args1 = length args2 then fold2 do_term_pair args1 args2 else K NONE) end | do_term_pair (ATerm ((s1, _), args1)) (ATerm ((s2, _), args2)) (SOME subst) = (case apply2 is_tptp_variable (s1, s2) of (true, true) => (case AList.lookup (op =) subst s1 of SOME s2' => if s2' = s2 then SOME subst else NONE | NONE => if null (AList.find (op =) subst s2) then SOME ((s1, s2) :: subst) else NONE) | (false, false) => if s1 = s2 then SOME subst else NONE | _ => NONE) |> (if length args1 = length args2 then fold2 do_term_pair args1 args2 else K NONE) | do_term_pair _ _ _ = NONE in SOME subst |> do_term_pair tm1 tm2 |> is_some end fun is_same_formula comm subst (AQuant (q1, xs1, phi1)) (AQuant (q2, xs2, phi2)) = q1 = q2 andalso length xs1 = length xs2 andalso is_same_formula comm ((map fst xs1 ~~ map fst xs2) @ subst) phi1 phi2 | is_same_formula comm subst (AConn (c1, phis1)) (AConn (c2, phis2)) = c1 = c2 andalso length phis1 = length phis2 andalso forall (uncurry (is_same_formula comm subst)) (phis1 ~~ phis2) | is_same_formula comm subst (AAtom (tm1 as ATerm (("equal", tys), [tm11, tm12]))) (AAtom tm2) = is_same_term subst tm1 tm2 orelse (comm andalso is_same_term subst (ATerm (("equal", tys), [tm12, tm11])) tm2) | is_same_formula _ subst (AAtom tm1) (AAtom tm2) = is_same_term subst tm1 tm2 | is_same_formula _ _ _ _ = false fun matching_formula_line_identifier phi (Formula ((ident, _), _, phi', _, _)) = if is_same_formula true [] phi phi' then SOME (ident, phi') else NONE | matching_formula_line_identifier _ _ = NONE fun find_formula_in_problem phi = maps snd #> map_filter (matching_formula_line_identifier phi) #> try (single o hd) #> the_default [] fun commute_eq (AAtom (ATerm ((s, tys), tms))) = AAtom (ATerm ((s, tys), rev tms)) | commute_eq _ = raise Fail "expected equation" fun role_of_tptp_string "axiom" = Axiom | role_of_tptp_string "definition" = Definition | role_of_tptp_string "lemma" = Lemma | role_of_tptp_string "hypothesis" = Hypothesis | role_of_tptp_string "conjecture" = Conjecture | role_of_tptp_string "negated_conjecture" = Negated_Conjecture | role_of_tptp_string "plain" = Plain | role_of_tptp_string "type" = Type_Role | role_of_tptp_string _ = Unknown fun parse_one_in_list xs = foldl1 (op ||) (map Scan.this_string xs) val tptp_literal_binary_ops = [tptp_equal, tptp_not_equal] val tptp_nonliteral_binary_ops = [tptp_and, tptp_or, tptp_implies, tptp_iff, tptp_if, tptp_not_and, tptp_not_or, tptp_not_iff] fun parse_literal_binary_op x = (parse_one_in_list tptp_literal_binary_ops >> (fn c => if c = tptp_equal then "equal" else c)) x fun parse_nonliteral_binary_op x = (parse_one_in_list tptp_nonliteral_binary_ops >> (fn c => if c = tptp_equal then "equal" else c)) x val parse_fol_quantifier = parse_one_in_list [tptp_forall, tptp_exists, tptp_lambda, tptp_hilbert_choice, tptp_hilbert_the] val parse_hol_quantifier = parse_one_in_list [tptp_ho_forall, tptp_ho_exists, tptp_hilbert_choice, tptp_hilbert_the] fun mk_ho_of_fo_quant q = if q = tptp_forall then tptp_ho_forall else if q = tptp_exists then tptp_ho_exists else if q = tptp_hilbert_choice then tptp_hilbert_choice else if q = tptp_hilbert_the then tptp_hilbert_the else raise Fail ("unrecognized quantification: " ^ q) fun remove_hol_app (ATerm ((s, ty), args)) = if s = tptp_app then (case args of ATerm (f, xs) :: ys => remove_hol_app (ATerm (f, xs @ ys)) | AAbs ((var, phi), xs) :: ys => remove_hol_app (AAbs ((var, phi), xs @ ys))) else ATerm ((s, ty), map remove_hol_app args) | remove_hol_app (AAbs ((var, phi), args)) = AAbs ((var, remove_hol_app phi), map remove_hol_app args) fun parse_hol_typed_var x = (Scan.repeat (scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type) --| Scan.option (Scan.this_string ",")) || $$ "(" |-- parse_hol_typed_var --| $$ ")") x fun parse_simple_hol_term x = (parse_fol_quantifier -- ($$ "[" |-- parse_hol_typed_var --| $$ "]" --| $$ ":") -- parse_hol_term >> (fn ((q, ys), t) => fold_rev (fn (var, ty) => fn r => AAbs (((var, the_default dummy_atype ty), r), []) |> (if tptp_lambda <> q then mk_app (q |> mk_ho_of_fo_quant |> mk_simple_aterm) else I)) ys t) || Scan.this_string tptp_not |-- parse_hol_term >> mk_app (mk_simple_aterm tptp_not) || scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type) >> (fn (var, typ_opt) => ATerm ((var, the_list typ_opt), [])) || parse_hol_quantifier >> mk_simple_aterm || $$ "(" |-- parse_hol_term --| $$ ")" || parse_literal_binary_op >> mk_simple_aterm || parse_nonliteral_binary_op >> mk_simple_aterm) x and parse_applied_hol_term x = (parse_simple_hol_term -- Scan.repeat (Scan.this_string tptp_app |-- parse_simple_hol_term) >> (fn (t1, tis) => fold (fn ti => fn left => mk_app left ti) tis t1)) x and parse_literal_hol_term x = (parse_applied_hol_term -- Scan.repeat (parse_literal_binary_op -- parse_applied_hol_term) >> (fn (t1, c_ti_s) => fold (fn (c, ti) => fn left => mk_apps (mk_simple_aterm c) [left, ti]) c_ti_s t1)) x and parse_hol_term x = (parse_literal_hol_term -- Scan.repeat (parse_nonliteral_binary_op -- parse_literal_hol_term) >> (fn (t1, c_ti_s) => fold (fn (c, ti) => fn left => mk_apps (mk_simple_aterm c) [left, ti]) c_ti_s t1)) x fun parse_hol_formula x = (parse_hol_term #>> remove_hol_app #>> AAtom) x -fun parse_tstp_hol_line problem = +fun parse_tstp_hol_line full problem = (Scan.this_string tptp_thf -- $$ "(") |-- scan_general_id --| $$ "," - -- Symbol.scan_ascii_id --| $$ "," -- (parse_hol_formula || skip_term >> K dummy_phi) + -- Symbol.scan_ascii_id --| $$ "," -- + (if full then parse_hol_formula || skip_term >> K dummy_phi else skip_term >> K dummy_phi) -- parse_tstp_extra_arguments --| $$ ")" --| $$ "." >> (fn (((num, role), phi), src) => let val role' = role_of_tptp_string role val ((name, phi), rule, deps) = (case src of SOME (File_Source (_, SOME s)) => if role' = Definition then (((num, map fst (find_formula_in_problem phi problem)), phi), "", []) else (((num, [s]), phi), "", []) | SOME (Inference_Source (rule, deps)) => (((num, []), phi), rule, deps) | SOME (Introduced_Source rule) => (((num, []), phi), rule, []) | _ => (((num, [num]), phi), "", [])) in [(name, role', phi, rule, map (rpair []) deps)] end) -fun parse_tstp_fol_line problem = +fun parse_tstp_fol_line full problem = ((Scan.this_string tptp_cnf || Scan.this_string tptp_fof || Scan.this_string tptp_tff) -- $$ "(") |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ "," - -- (parse_fol_formula || skip_term >> K dummy_phi) -- parse_tstp_extra_arguments - --| $$ ")" --| $$ "." + -- (if full then parse_fol_formula || skip_term >> K dummy_phi else skip_term >> K dummy_phi) + -- parse_tstp_extra_arguments --| $$ ")" --| $$ "." >> (fn (((num, role0), phi), src) => let val role = role_of_tptp_string role0 val ((name, phi), role', rule, deps) = (* Waldmeister isn't exactly helping. *) (case src of SOME (File_Source (_, SOME s)) => (if s = waldmeister_conjecture_name then (case find_formula_in_problem (mk_anot phi) problem of (* Waldmeister hack: Get the original orientation of the equation to avoid confusing Isar. *) [(s, phi')] => ((num, [s]), phi |> not (is_same_formula false [] (mk_anot phi) phi') ? commute_eq) | _ => ((num, []), phi)) else ((num, [s]), phi), role, "", []) | SOME (File_Source _) => (((num, map fst (find_formula_in_problem phi problem)), phi), role, "", []) | SOME (Inference_Source (rule, deps)) => (((num, []), phi), role, rule, deps) | SOME (Introduced_Source rule) => (((num, []), phi), Lemma, rule, []) | _ => (((num, [num]), phi), role, "", [])) fun mk_step () = (name, role', phi, rule, map (rpair []) deps) in [(case role' of Definition => (case phi of AAtom (ATerm (("equal", _), _)) => (* Vampire's equality proxy axiom *) (name, Definition, phi, rule, map (rpair []) deps) | _ => mk_step ()) | _ => mk_step ())] end) -fun parse_tstp_line problem = parse_tstp_fol_line problem || parse_tstp_hol_line problem +fun parse_tstp_line full problem = + parse_tstp_fol_line full problem + || parse_tstp_hol_line full problem (**** PARSING OF SPASS OUTPUT ****) (* SPASS returns clause references of the form "x.y". We ignore "y". *) val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id val parse_spass_annotations = Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name --| Scan.option ($$ ","))) [] (* We ignore the stars and the pluses that follow literals. *) fun parse_decorated_atom x = (parse_fol_atom --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")) x fun mk_horn ([], []) = AAtom (ATerm (("c_False", []), [])) | mk_horn (neg_lits, pos_lits) = foldr1 (uncurry (mk_aconn AOr)) (map mk_anot neg_lits @ pos_lits) fun parse_horn_clause x = (Scan.repeat parse_decorated_atom --| $$ "|" --| $$ "|" -- Scan.repeat parse_decorated_atom --| $$ "-" --| $$ ">" -- Scan.repeat parse_decorated_atom >> (mk_horn o apfst (op @))) x val parse_spass_debug = Scan.option ($$ "(" |-- Scan.repeat (scan_general_id --| Scan.option ($$ ",")) --| $$ ")") (* Syntax: [0:] || -> . derived from formulae * *) fun parse_spass_line x = (parse_spass_debug |-- scan_general_id --| $$ "[" --| Scan.many1 Symbol.is_digit --| $$ ":" -- Symbol.scan_ascii_id -- parse_spass_annotations --| $$ "]" -- parse_horn_clause --| $$ "." -- Scan.option (Scan.this_string "derived from formulae " |-- Scan.repeat (scan_general_id --| Scan.option ($$ " "))) >> (fn ((((num, rule), deps), u), names) => [((num, these names), Unknown, u, rule, map (rpair []) deps)])) x fun core_inference inf fact = ((fact, [fact]), Unknown, dummy_phi, inf, []) -fun parse_line local_name problem = +fun parse_line full local_name problem = (* Satallax is handled separately, in "atp_satallax.ML". *) if local_name = spassN then parse_spass_line - else parse_tstp_line problem + else parse_tstp_line full problem fun core_of_agsyhol_proof s = (case split_lines s of "The transformed problem consists of the following conjectures:" :: conj :: _ :: proof_term :: _ => SOME (unprefix " " conj :: find_enclosed "<<" ">>" proof_term) | _ => NONE) fun clean_up_dependencies _ [] = [] | clean_up_dependencies seen ((name, role, u, rule, deps) :: steps) = (name, role, u, rule, map_filter (fn dep => find_first (is_same_atp_step dep) seen) deps) :: clean_up_dependencies (name :: seen) steps fun clean_up_atp_proof_dependencies proof = clean_up_dependencies [] proof fun map_term_names_in_atp_proof f = let fun map_type (AType ((s, clss), tys)) = AType ((f s, map f clss), map map_type tys) | map_type (AFun (ty, ty')) = AFun (map_type ty, map_type ty') | map_type (APi (ss, ty)) = APi (map f ss, map_type ty) fun map_term (ATerm ((s, tys), ts)) = ATerm ((f s, map map_type tys), map map_term ts) | map_term (AAbs (((s, ty), tm), args)) = AAbs (((f s, map_type ty), map_term tm), map map_term args) fun map_formula (AQuant (q, xs, phi)) = AQuant (q, map (apfst f) xs, map_formula phi) | map_formula (AConn (c, phis)) = AConn (c, map map_formula phis) | map_formula (AAtom t) = AAtom (map_term t) fun map_step (name, role, phi, rule, deps) = (name, role, map_formula phi, rule, deps) in map map_step end fun nasty_name pool s = Symtab.lookup pool s |> the_default s fun nasty_atp_proof pool = not (Symtab.is_empty pool) ? map_term_names_in_atp_proof (nasty_name pool) fun string_of_list f xs = enclose "[" "]" (commas (map f xs)) fun string_of_atp_step_name (s, ss) = "(" ^ s ^ ", " ^ string_of_list I ss ^ ")" fun string_of_atp_step f g (name, role, x, y, names) = let val name' = string_of_atp_step_name name val role' = ATP_Problem.tptp_string_of_role role val x' = f x val y' = g y val names' = string_of_list string_of_atp_step_name names in "(" ^ name' ^ ", " ^ role' ^ ", " ^ x' ^ ", " ^ y' ^ ", " ^ names' ^ ")" end -fun parse_proof local_name problem = +fun parse_proof full local_name problem = strip_spaces_except_between_idents #> raw_explode #> Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ()) - (Scan.finite Symbol.stopper (Scan.repeats1 (parse_line local_name problem)))) - #> fst + (Scan.finite Symbol.stopper (Scan.repeats1 (parse_line full local_name problem)))) + #> (fn (proof, ss) => if null ss then proof else raise UNRECOGNIZED_ATP_PROOF ()) -fun atp_proof_of_tstplike_proof _ _ "" = [] - | atp_proof_of_tstplike_proof local_prover problem tstp = +fun atp_proof_of_tstplike_proof _ _ _ "" = [] + | atp_proof_of_tstplike_proof full local_prover problem tstp = (case core_of_agsyhol_proof tstp of SOME facts => facts |> map (core_inference agsyhol_core_rule) | NONE => - tstp ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *) - |> parse_proof local_prover problem + tstp + |> parse_proof full local_prover problem |> local_prover = vampireN ? perhaps (try (sort (vampire_step_name_ord o apply2 #1)))) end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML @@ -1,408 +1,400 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_commands.ML Author: Jasmin Blanchette, TU Muenchen Adds "sledgehammer" and related commands to Isabelle/Isar's outer syntax. *) signature SLEDGEHAMMER_COMMANDS = sig type params = Sledgehammer_Prover.params val provers : string Unsynchronized.ref val default_params : theory -> (string * string) list -> params val parse_params: (string * string) list parser end; structure Sledgehammer_Commands : SLEDGEHAMMER_COMMANDS = struct open ATP_Util open ATP_Problem_Generate open ATP_Proof open ATP_Proof_Reconstruct open Sledgehammer_Util open Sledgehammer_Fact open Sledgehammer_ATP_Systems open Sledgehammer_Prover open Sledgehammer_Prover_SMT open Sledgehammer_Prover_Minimize open Sledgehammer_MaSh open Sledgehammer val runN = "run" val supported_proversN = "supported_provers" val refresh_tptpN = "refresh_tptp" (** Sledgehammer commands **) fun add_fact_override ns : fact_override = {add = ns, del = [], only = false} fun del_fact_override ns : fact_override = {add = [], del = ns, only = false} fun only_fact_override ns : fact_override = {add = ns, del = [], only = true} fun merge_fact_override_pairwise (r1 : fact_override) (r2 : fact_override) = {add = #add r1 @ #add r2, del = #del r1 @ #del r2, only = #only r1 andalso #only r2} fun merge_fact_overrides rs = fold merge_fact_override_pairwise rs (only_fact_override []) (*** parameters ***) val provers = Unsynchronized.ref "" type raw_param = string * string list val default_default_params = [("debug", "false"), ("verbose", "false"), ("overlord", "false"), ("spy", "false"), ("type_enc", "smart"), ("strict", "false"), ("lam_trans", "smart"), ("uncurried_aliases", "smart"), ("learn", "true"), ("fact_filter", "smart"), ("induction_rules", "smart"), ("max_facts", "smart"), ("fact_thresholds", "0.45 0.85"), ("max_mono_iters", "smart"), ("max_new_mono_instances", "smart"), ("max_proofs", "4"), ("isar_proofs", "smart"), ("compress", "smart"), ("try0", "true"), ("smt_proofs", "true"), ("minimize", "true"), ("slices", string_of_int (6 * Multithreading.max_threads ())), ("preplay_timeout", "1")] val alias_params = [("prover", ("provers", [])), (* undocumented *) ("dont_preplay", ("preplay_timeout", ["0"])), ("dont_compress", ("compress", ["1"])), ("dont_slice", ("slices", ["1"])), ("isar_proof", ("isar_proofs", [])) (* legacy *)] val negated_alias_params = [("no_debug", "debug"), ("quiet", "verbose"), ("no_overlord", "overlord"), ("dont_spy", "spy"), ("non_strict", "strict"), ("no_uncurried_aliases", "uncurried_aliases"), ("dont_learn", "learn"), ("no_isar_proofs", "isar_proofs"), ("dont_minimize", "minimize"), ("dont_try0", "try0"), ("no_smt_proofs", "smt_proofs")] val property_dependent_params = ["provers", "timeout"] fun is_known_raw_param s = AList.defined (op =) default_default_params s orelse AList.defined (op =) alias_params s orelse AList.defined (op =) negated_alias_params s orelse member (op =) property_dependent_params s orelse s = "expect" fun is_prover_list ctxt s = (case space_explode " " s of ss as _ :: _ => forall (is_prover_supported ctxt) ss | _ => false) fun unalias_raw_param (name, value) = (case AList.lookup (op =) alias_params name of SOME (name', []) => (name', value) | SOME (param' as (name', _)) => if value <> ["false"] then param' else error ("Parameter " ^ quote name ^ " cannot be set to \"false\" (cf. " ^ quote name' ^ ")") | NONE => (case AList.lookup (op =) negated_alias_params name of SOME name' => (name', (case value of ["false"] => ["true"] | ["true"] => ["false"] | [] => ["false"] | _ => value)) | NONE => (name, value))) val any_type_enc = type_enc_of_string Strict "erased" (* "provers =", "type_enc =", "lam_trans =", "fact_filter =", and "max_facts =" can be omitted. For the last four, this is a secret feature. *) fun normalize_raw_param ctxt = unalias_raw_param #> (fn (name, value) => if is_known_raw_param name then (name, value) else if null value then if is_prover_list ctxt name then ("provers", [name]) else if can (type_enc_of_string Strict) name then ("type_enc", [name]) else if can (trans_lams_of_string ctxt any_type_enc) name then ("lam_trans", [name]) else if member (op =) fact_filters name then ("fact_filter", [name]) else if is_some (Int.fromString name) then ("max_facts", [name]) else error ("Unknown parameter: " ^ quote name) else error ("Unknown parameter: " ^ quote name)) (* Ensures that type encodings such as "mono_native?" and "poly_guards!!" are read correctly. *) val implode_param = strip_spaces_except_between_idents o space_implode " " (* FIXME: use "Generic_Data" *) structure Data = Theory_Data ( type T = raw_param list val empty = default_default_params |> map (apsnd single) fun merge data : T = AList.merge (op =) (K true) data ) -fun remotify_prover_if_supported_and_not_already_remote ctxt name = - if String.isPrefix remote_prefix name then - SOME name - else - let val remote_name = remote_prefix ^ name in - if is_prover_supported ctxt remote_name then SOME remote_name else NONE - end - (* The first ATP of the list is used by Auto Sledgehammer. *) fun default_provers_param_value ctxt = filter (is_prover_installed ctxt) (smts ctxt @ atps) \ \see also \<^system_option>\sledgehammer_provers\\ |> implode_param fun set_default_raw_param param thy = let val ctxt = Proof_Context.init_global thy in thy |> Data.map (AList.update (op =) (normalize_raw_param ctxt param)) end fun default_raw_params thy = let val ctxt = Proof_Context.init_global thy in Data.get thy |> fold (AList.default (op =)) [("provers", [(case !provers of "" => default_provers_param_value ctxt | s => s)]), ("timeout", let val timeout = Options.default_int \<^system_option>\sledgehammer_timeout\ in [if timeout <= 0 then "none" else string_of_int timeout] end)] end fun extract_params mode default_params override_params = let val raw_params = rev override_params @ rev default_params val lookup = Option.map implode_param o AList.lookup (op =) raw_params val lookup_string = the_default "" o lookup fun general_lookup_bool option default_value name = (case lookup name of SOME s => parse_bool_option option name s | NONE => default_value) val lookup_bool = the o general_lookup_bool false (SOME false) fun lookup_time name = (case lookup name of SOME s => parse_time name s | NONE => Time.zeroTime) fun lookup_int name = (case lookup name of NONE => 0 | SOME s => (case Int.fromString s of SOME n => n | NONE => error ("Parameter " ^ quote name ^ " must be assigned an integer value"))) fun lookup_real name = (case lookup name of NONE => 0.0 | SOME s => (case Real.fromString s of SOME n => n | NONE => error ("Parameter " ^ quote name ^ " must be assigned a real value"))) fun lookup_real_pair name = (case lookup name of NONE => (0.0, 0.0) | SOME s => (case s |> space_explode " " |> map Real.fromString of [SOME r1, SOME r2] => (r1, r2) | _ => error ("Parameter " ^ quote name ^ " must be assigned a pair of floating-point \ \values (e.g., \"0.6 0.95\")"))) fun lookup_option lookup' name = (case lookup name of SOME "smart" => NONE | _ => SOME (lookup' name)) val debug = mode <> Auto_Try andalso lookup_bool "debug" val verbose = debug orelse (mode <> Auto_Try andalso lookup_bool "verbose") val overlord = lookup_bool "overlord" val spy = getenv "SLEDGEHAMMER_SPY" = "yes" orelse lookup_bool "spy" val provers = lookup_string "provers" |> space_explode " " |> mode = Auto_Try ? single o hd val type_enc = if mode = Auto_Try then NONE else (case lookup_string "type_enc" of "smart" => NONE | s => (type_enc_of_string Strict s; SOME s)) val strict = mode = Auto_Try orelse lookup_bool "strict" val lam_trans = lookup_option lookup_string "lam_trans" val uncurried_aliases = lookup_option lookup_bool "uncurried_aliases" val learn = lookup_bool "learn" val fact_filter = lookup_option lookup_string "fact_filter" |> mode = Auto_Try ? (fn NONE => SOME mepoN | sf => sf) val induction_rules = lookup_option (the o induction_rules_of_string o lookup_string) "induction_rules" val max_facts = lookup_option lookup_int "max_facts" val fact_thresholds = lookup_real_pair "fact_thresholds" val max_mono_iters = lookup_option lookup_int "max_mono_iters" val max_new_mono_instances = lookup_option lookup_int "max_new_mono_instances" val max_proofs = lookup_int "max_proofs" val isar_proofs = lookup_option lookup_bool "isar_proofs" val compress = Option.map (curry Real.max 1.0) (lookup_option lookup_real "compress") val try0 = lookup_bool "try0" val smt_proofs = lookup_bool "smt_proofs" val minimize = mode <> Auto_Try andalso lookup_bool "minimize" val slices = if mode = Auto_Try then 1 else Int.max (1, lookup_int "slices") val timeout = lookup_time "timeout" val preplay_timeout = lookup_time "preplay_timeout" val expect = lookup_string "expect" in {debug = debug, verbose = verbose, overlord = overlord, spy = spy, provers = provers, type_enc = type_enc, strict = strict, lam_trans = lam_trans, uncurried_aliases = uncurried_aliases, learn = learn, fact_filter = fact_filter, induction_rules = induction_rules, max_facts = max_facts, fact_thresholds = fact_thresholds, max_mono_iters = max_mono_iters, max_new_mono_instances = max_new_mono_instances, max_proofs = max_proofs, isar_proofs = isar_proofs, compress = compress, try0 = try0, smt_proofs = smt_proofs, minimize = minimize, slices = slices, timeout = timeout, preplay_timeout = preplay_timeout, expect = expect} end fun get_params mode = extract_params mode o default_raw_params fun default_params thy = get_params Normal thy o map (apsnd single) val silence_state = Proof.map_contexts (Try0.silence_methods false #> Config.put SMT_Config.verbose false) (* Sledgehammer the given subgoal *) val default_learn_prover_timeout = 2.0 fun hammer_away override_params writeln_result subcommand opt_i fact_override state0 = let (* We generally want chained facts to be picked up by the relevance filter, because it can then give it a proper name, which is useful for a variety of reasons (minimization, Isar proofs, verbose output, machine learning). However, if the fact is available by no other means (not even backticks), as "f" would be in "using f unfolding f'" after unfolding, we have no choice but to insert it into the state as an additional hypothesis. *) val {facts = chained0, ...} = Proof.goal state0 val (inserts, keep_chained) = if null chained0 orelse #only fact_override then (chained0, []) else let val ctxt0 = Proof.context_of state0 in List.partition (is_useful_unnamed_local_fact ctxt0) chained0 end val state = state0 |> (if null inserts then I else Proof.refine_insert inserts #> Proof.set_facts keep_chained) |> silence_state val thy = Proof.theory_of state val ctxt = Proof.context_of state val override_params = override_params |> map (normalize_raw_param ctxt) in if subcommand = runN then let val i = the_default 1 opt_i in ignore (run_sledgehammer (get_params Normal thy override_params) Normal writeln_result i fact_override state) end else if subcommand = supported_proversN then supported_provers ctxt else if subcommand = unlearnN then mash_unlearn ctxt else if subcommand = learn_isarN orelse subcommand = learn_proverN orelse subcommand = relearn_isarN orelse subcommand = relearn_proverN then (if subcommand = relearn_isarN orelse subcommand = relearn_proverN then mash_unlearn ctxt else (); mash_learn ctxt (* TODO: Use MaSh mode instead and have the special defaults hardcoded in "get_params" *) (get_params Normal thy ([("timeout", [string_of_real default_learn_prover_timeout]), ("slice", ["false"])] @ override_params @ [("preplay_timeout", ["0"])])) fact_override (#facts (Proof.goal state)) (subcommand = learn_proverN orelse subcommand = relearn_proverN)) else if subcommand = refresh_tptpN then refresh_systems_on_tptp () else error ("Unknown subcommand: " ^ quote subcommand) end fun string_of_raw_param (key, values) = key ^ (case implode_param values of "" => "" | value => " = " ^ value) val parse_query_bang = \<^keyword>\?\ || \<^keyword>\!\ || \<^keyword>\!!\ val parse_key = Scan.repeat1 (Parse.embedded || parse_query_bang) >> implode_param val parse_value = Scan.repeat1 (Parse.name || Parse.float_number || parse_query_bang) val parse_param = parse_key -- Scan.optional (\<^keyword>\=\ |-- parse_value) [] val parse_raw_params = Scan.optional (Args.bracks (Parse.list parse_param)) [] val parse_params = parse_raw_params >> map (apsnd implode_param) val parse_fact_refs = Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse.thm) val parse_fact_override_chunk = (Args.add |-- Args.colon |-- parse_fact_refs >> add_fact_override) || (Args.del |-- Args.colon |-- parse_fact_refs >> del_fact_override) || (parse_fact_refs >> only_fact_override) val parse_fact_override = Scan.optional (Args.parens (Scan.repeat parse_fact_override_chunk >> merge_fact_overrides)) no_fact_override val _ = Outer_Syntax.command \<^command_keyword>\sledgehammer\ "search for first-order proof using automatic theorem provers" (Scan.optional Parse.name runN -- parse_raw_params -- parse_fact_override -- Scan.option Parse.nat >> (fn (((subcommand, params), fact_override), opt_i) => Toplevel.keep_proof (hammer_away params NONE subcommand opt_i fact_override o Toplevel.proof_of))) val _ = Outer_Syntax.command \<^command_keyword>\sledgehammer_params\ "set and display the default parameters for Sledgehammer" (parse_raw_params >> (fn params => Toplevel.theory (fold set_default_raw_param params #> tap (fn thy => writeln ("Default parameters for Sledgehammer:\n" ^ (case rev (default_raw_params thy) of [] => "none" | params => params |> map string_of_raw_param |> sort_strings |> cat_lines)))))) val _ = Try.tool_setup {name = sledgehammerN, weight = 40, auto_option = \<^system_option>\auto_sledgehammer\, body = fn auto => fn state => let val thy = Proof.theory_of state val mode = if auto then Auto_Try else Try val i = 1 in run_sledgehammer (get_params mode thy []) mode NONE i no_fact_override (silence_state state) |> apsnd (map_prod short_string_of_sledgehammer_outcome single) end} val _ = Query_Operation.register {name = sledgehammerN, pri = 0} (fn {state = st, args, writeln_result, ...} => (case try Toplevel.proof_of st of SOME state => let val [provers_arg, isar_proofs_arg, try0_arg] = args val override_params = ((if provers_arg = "" then [] else [("provers", space_explode " " provers_arg)]) @ [("isar_proofs", [if isar_proofs_arg = "true" then "true" else "smart"]), ("try0", [try0_arg]), ("debug", ["false"]), ("verbose", ["false"]), ("overlord", ["false"])]); in hammer_away override_params (SOME writeln_result) runN NONE no_fact_override state end | NONE => error "Unknown proof context")) end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML @@ -1,349 +1,358 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML Author: Fabian Immler, TU Muenchen Author: Makarius Author: Jasmin Blanchette, TU Muenchen ATPs as Sledgehammer provers. *) signature SLEDGEHAMMER_PROVER_ATP = sig type mode = Sledgehammer_Prover.mode type prover = Sledgehammer_Prover.prover val atp_problem_dest_dir : string Config.T val atp_proof_dest_dir : string Config.T val atp_problem_prefix : string Config.T val atp_completish : int Config.T val atp_full_names : bool Config.T val run_atp : mode -> string -> prover end; structure Sledgehammer_Prover_ATP : SLEDGEHAMMER_PROVER_ATP = struct open ATP_Util open ATP_Problem open ATP_Problem_Generate open ATP_Proof open ATP_Proof_Reconstruct open Sledgehammer_Util open Sledgehammer_Proof_Methods open Sledgehammer_Isar open Sledgehammer_ATP_Systems open Sledgehammer_Prover (* Empty string means create files in Isabelle's temporary files directory. *) val atp_problem_dest_dir = Attrib.setup_config_string \<^binding>\sledgehammer_atp_problem_dest_dir\ (K "") val atp_proof_dest_dir = Attrib.setup_config_string \<^binding>\sledgehammer_atp_proof_dest_dir\ (K "") val atp_problem_prefix = Attrib.setup_config_string \<^binding>\sledgehammer_atp_problem_prefix\ (K "prob") val atp_completish = Attrib.setup_config_int \<^binding>\sledgehammer_atp_completish\ (K 0) (* In addition to being easier to read, readable names are often much shorter, especially if types are mangled in names. This makes a difference for some provers (e.g., E). For these reason, short names are enabled by default. *) val atp_full_names = Attrib.setup_config_bool \<^binding>\sledgehammer_atp_full_names\ (K false) fun choose_type_enc strictness format good_type_enc = type_enc_of_string strictness good_type_enc |> adjust_type_enc format fun has_bound_or_var_of_type pred = exists_subterm (fn Var (_, T as Type _) => pred T | Abs (_, T as Type _, _) => pred T | _ => false) (* Unwanted equalities are those between a (bound or schematic) variable that does not properly occur in the second operand. *) val is_exhaustive_finite = let fun is_bad_equal (Var z) t = not (exists_subterm (fn Var z' => z = z' | _ => false) t) | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j)) | is_bad_equal _ _ = false fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1 fun do_formula pos t = (case (pos, t) of (_, \<^Const_>\Trueprop for t1\) => do_formula pos t1 | (true, Const (\<^const_name>\Pure.all\, _) $ Abs (_, _, t')) => do_formula pos t' | (true, Const (\<^const_name>\All\, _) $ Abs (_, _, t')) => do_formula pos t' | (false, Const (\<^const_name>\Ex\, _) $ Abs (_, _, t')) => do_formula pos t' | (_, \<^Const_>\Pure.imp for t1 t2\) => do_formula (not pos) t1 andalso (t2 = \<^prop>\False\ orelse do_formula pos t2) | (_, \<^Const_>\implies for t1 t2\) => do_formula (not pos) t1 andalso (t2 = \<^Const>\False\ orelse do_formula pos t2) | (_, \<^Const_>\Not for t1\) => do_formula (not pos) t1 | (true, \<^Const_>\disj for t1 t2\) => forall (do_formula pos) [t1, t2] | (false, \<^Const_>\conj for t1 t2\) => forall (do_formula pos) [t1, t2] | (true, Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) => do_equals t1 t2 | (true, Const (\<^const_name>\Pure.eq\, _) $ t1 $ t2) => do_equals t1 t2 | _ => false) in do_formula true end (* Facts containing variables of finite types such as "unit" or "bool" or of the form "ALL x. x = A | x = B | x = C" are likely to lead to untypable proofs for unsound type encodings. *) fun is_dangerous_prop ctxt = transform_elim_prop #> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf is_exhaustive_finite) val mono_max_privileged_facts = 10 fun suffix_of_mode Auto_Try = "_try" | suffix_of_mode Try = "_try" | suffix_of_mode Normal = "" | suffix_of_mode MaSh = "" | suffix_of_mode Minimize = "_min" (* Important messages are important but not so important that users want to see them each time. *) val atp_important_message_keep_quotient = 25 fun run_atp mode name ({debug, verbose, overlord, strict, fact_filter, max_facts, max_mono_iters, max_new_mono_instances, isar_proofs, compress, try0, smt_proofs, minimize, slices, timeout, preplay_timeout, spy, ...} : params) ({comment, state, goal, subgoal, subgoal_count, factss, found_proof} : prover_problem) slice = let val thy = Proof.theory_of state val ctxt = Proof.context_of state val ((good_max_facts, good_fact_filter), ATP_Slice (good_format, good_type_enc, good_lam_trans, good_uncurried_aliases, extra)) = slice val {exec, arguments, proof_delims, known_failures, prem_role, good_max_mono_iters, good_max_new_mono_instances, ...} = get_atp thy name () val full_proofs = isar_proofs |> the_default (mode = Minimize) val local_name = perhaps (try (unprefix remote_prefix)) name val completish = Config.get ctxt atp_completish val atp_mode = if completish > 0 then Sledgehammer_Completish completish else Sledgehammer val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt val (problem_dest_dir, proof_dest_dir, problem_prefix) = if overlord then overlord_file_location_of_prover name |> (fn (dir, prefix) => (dir, dir, prefix)) else (Config.get ctxt atp_problem_dest_dir, Config.get ctxt atp_proof_dest_dir, Config.get ctxt atp_problem_prefix) val problem_file_name = Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^ suffix_of_mode mode ^ "_" ^ string_of_int subgoal) |> Path.ext "p" val prob_path = if problem_dest_dir = "" then File.tmp_path problem_file_name else if File.exists (Path.explode problem_dest_dir) then Path.explode problem_dest_dir + problem_file_name else error ("No such directory: " ^ quote problem_dest_dir) val executable = (case find_first (fn var => getenv var <> "") (fst exec) of SOME var => let val pref = getenv var ^ "/" val paths = map (Path.explode o prefix pref) (if ML_System.platform_is_windows then map (suffix ".exe") (snd exec) @ snd exec else snd exec); in (case find_first File.exists paths of SOME path => path | NONE => error ("Bad executable: " ^ Path.print (hd paths))) end | NONE => error ("The environment variable " ^ quote (List.last (fst exec)) ^ " is not set")) fun run () = let fun monomorphize_facts facts = let val ctxt = ctxt |> repair_monomorph_context max_mono_iters good_max_mono_iters max_new_mono_instances good_max_new_mono_instances (* pseudo-theorem involving the same constants as the subgoal *) val subgoal_th = Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy val rths = facts |> chop mono_max_privileged_facts |>> map (pair 1 o snd) ||> map (pair 2 o snd) |> op @ |> cons (0, subgoal_th) in Monomorph.monomorph atp_schematic_consts_of ctxt rths |> tl |> curry ListPair.zip (map fst facts) |> maps (fn (name, rths) => map (pair name o zero_var_indexes o snd) rths) end val effective_fact_filter = fact_filter |> the_default good_fact_filter val facts = facts_of_filter effective_fact_filter factss val num_facts = (case max_facts of NONE => good_max_facts | SOME max_facts => max_facts) |> Integer.min (length facts) val strictness = if strict then Strict else Non_Strict val type_enc = choose_type_enc strictness good_format good_type_enc val run_timeout = slice_timeout slices timeout val generous_run_timeout = if mode = MaSh then one_day else run_timeout val ({elapsed, ...}, atp_problem_data as (atp_problem, _, _, _)) = Timing.timing (fn () => let val sound = is_type_enc_sound type_enc val generate_info = (case good_format of DFG _ => true | _ => false) val readable_names = not (Config.get ctxt atp_full_names) in facts |> not sound ? filter_out (is_dangerous_prop ctxt o Thm.prop_of o snd) |> take num_facts |> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts |> map (apsnd Thm.prop_of) |> generate_atp_problem ctxt generate_info good_format prem_role type_enc atp_mode good_lam_trans good_uncurried_aliases readable_names true hyp_ts concl_t end) () val () = spying spy (fn () => (state, subgoal, name, "Generating ATP problem in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms")) fun sel_weights () = atp_problem_selection_weights atp_problem fun ord_info () = atp_problem_term_order_info atp_problem val ord = effective_term_order ctxt name val args = arguments ctxt full_proofs extra run_timeout prob_path (ord, ord_info, sel_weights) val command = space_implode " " (File.bash_path executable :: args) fun run_command () = if exec = isabelle_scala_function then let val {output, timing} = SystemOnTPTP.run_system_encoded args in (output, timing) end else let val res = Isabelle_System.bash_process (Bash.script command |> Bash.redirect) in (Process_Result.out res, Process_Result.timing_elapsed res) end val _ = atp_problem |> lines_of_atp_problem good_format ord ord_info |> (exec <> isabelle_scala_function) ? cons ("% " ^ command ^ "\n" ^ (if comment = "" then "" else "% " ^ comment ^ "\n")) |> File.write_list prob_path - val ((output, run_time), (atp_proof, outcome)) = + val ((output, run_time), ((atp_proof, tstplike_proof), outcome)) = Timeout.apply generous_run_timeout run_command () |>> overlord ? (fn output => prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") output) |> (fn accum as (output, _) => (accum, extract_tstplike_proof_and_outcome verbose proof_delims known_failures output - |>> atp_proof_of_tstplike_proof (perhaps (try (unprefix remote_prefix)) name) - atp_problem - handle UNRECOGNIZED_ATP_PROOF () => ([], SOME ProofUnparsable))) - handle Timeout.TIMEOUT _ => (("", run_timeout), ([], SOME TimedOut)) - | ERROR msg => (("", Time.zeroTime), ([], SOME (UnknownError msg))) + |>> `(atp_proof_of_tstplike_proof false (perhaps (try (unprefix remote_prefix)) name) + atp_problem) + handle UNRECOGNIZED_ATP_PROOF () => (([], ""), SOME ProofUnparsable))) + handle Timeout.TIMEOUT _ => (("", run_timeout), (([], ""), SOME TimedOut)) + | ERROR msg => (("", Time.zeroTime), (([], ""), SOME (UnknownError msg))) val () = spying spy (fn () => (state, subgoal, name, "Running command in " ^ string_of_int (Time.toMilliseconds run_time) ^ " ms")) val outcome = (case outcome of NONE => (case used_facts_in_unsound_atp_proof ctxt (map fst facts) atp_proof of SOME facts => let val failure = UnsoundProof (is_type_enc_sound type_enc, sort string_ord facts) in if debug then (warning (string_of_atp_failure failure); NONE) else SOME failure end | NONE => (found_proof name; NONE)) | _ => outcome) in - (atp_problem_data, (output, run_time, facts, atp_proof, outcome), (good_format, type_enc)) + (atp_problem_data, + (output, run_time, facts, atp_problem, tstplike_proof, atp_proof, outcome), + (good_format, type_enc)) end (* If the problem file has not been exported, remove it; otherwise, export the proof file too. *) fun clean_up () = if problem_dest_dir = "" then (try File.rm prob_path; ()) else () - fun export (_, (output, _, _, _, _), _) = + fun export (_, (output, _, _, _, _, _, _), _) = let val proof_dest_dir_path = Path.explode proof_dest_dir val make_export_file_name = Path.split_ext #> apfst (Path.explode o suffix "_proof" o Path.implode) #> swap #> uncurry Path.ext in if proof_dest_dir = "" then Output.system_message "don't export proof" else if File.exists proof_dest_dir_path then File.write (proof_dest_dir_path + make_export_file_name problem_file_name) output else error ("No such directory: " ^ quote proof_dest_dir) end - val ((_, pool, lifted, sym_tab), (output, run_time, used_from, atp_proof, outcome), + val ((_, pool, lifted, sym_tab), + (output, run_time, used_from, atp_problem, tstplike_proof, atp_proof, outcome), (format, type_enc)) = with_cleanup clean_up run () |> tap export val important_message = if mode = Normal andalso Random.random_range 0 (atp_important_message_keep_quotient - 1) = 0 then extract_important_message output else "" val (used_facts, preferred_methss, message) = (case outcome of NONE => let val used_facts = sort_by fst (used_facts_in_atp_proof ctxt (map fst used_from) atp_proof) val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof val preferred_methss = (Metis_Method (NONE, NONE), bunches_of_proof_methods ctxt try0 smt_proofs needs_full_types - (if atp_proof_prefers_lifting atp_proof then liftingN else opaque_liftingN)) + (if atp_proof_prefers_lifting atp_proof then liftingN else opaque_liftingN)) in (used_facts, preferred_methss, fn preplay => let val _ = if verbose then writeln "Generating proof text..." else () fun isar_params () = let + val full_atp_proof = + atp_proof_of_tstplike_proof true (perhaps (try (unprefix remote_prefix)) name) + atp_problem tstplike_proof val metis_type_enc = - if is_typed_helper_used_in_atp_proof atp_proof then SOME full_typesN else NONE + if is_typed_helper_used_in_atp_proof full_atp_proof then + SOME full_typesN + else + NONE val metis_lam_trans = - if atp_proof_prefers_lifting atp_proof then SOME liftingN else NONE - val atp_proof = - atp_proof + if atp_proof_prefers_lifting full_atp_proof then SOME liftingN else NONE + val full_atp_proof = + full_atp_proof |> termify_atp_proof ctxt name format type_enc pool lifted sym_tab |> local_name = spassN ? introduce_spass_skolems |> factify_atp_proof (map fst used_from) hyp_ts concl_t in (verbose, (metis_type_enc, metis_lam_trans), preplay_timeout, compress, try0, - minimize, atp_proof, goal) + minimize, full_atp_proof, goal) end val one_line_params = (preplay (), proof_banner mode name, subgoal, subgoal_count) val num_chained = length (#facts (Proof.goal state)) in proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained one_line_params ^ (if important_message <> "" then "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message else "") end) end | SOME failure => ([], (Auto_Method (* dummy *), []), fn _ => string_of_atp_failure failure)) in {outcome = outcome, used_facts = used_facts, used_from = used_from, preferred_methss = preferred_methss, run_time = run_time, message = message} end end;