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,739 +1,739 @@ (* 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 zipperposition_define_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 : 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 : 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 *) val zipperposition_define_rule = "define" (* 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 = "Proof missing" | string_of_atp_failure ProofIncomplete = "Proof incomplete" | string_of_atp_failure ProofUnparsable = "Proof unparsable" | 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 = "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 | Define_Source 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_define_source x = (Scan.this_string "define" |-- $$ "(" |-- 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 || parse_define_source >> K Define_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 full problem = (Scan.this_string tptp_thf -- $$ "(") |-- scan_general_id --| $$ "," -- 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) => + >> (fn (((num, role0), phi), src) => let - val role' = role_of_tptp_string role - val ((name, phi), rule, deps) = + val role = role_of_tptp_string role0 + val (name, phi, role', rule, deps) = (case src of SOME (File_Source (_, SOME s)) => - if role' = Definition then - (((num, map fst (find_formula_in_problem phi problem)), phi), "", []) + if role = Definition then + ((num, map fst (find_formula_in_problem phi problem)), phi, role, "", []) else - (((num, [s]), phi), "", []) - | SOME (Inference_Source (rule, deps)) => (((num, []), phi), rule, deps) - | SOME (Introduced_Source rule) => (((num, []), phi), rule, []) - | SOME Define_Source => (((num, []), phi), zipperposition_define_rule, []) - | _ => (((num, [num]), phi), "", [])) + ((num, [s]), phi, role, "", []) + | SOME (Inference_Source (rule, deps)) => ((num, []), phi, role, rule, deps) + | SOME (Introduced_Source rule) => ((num, []), phi, Definition, rule, []) + | SOME Define_Source => ((num, []), phi, Definition, zipperposition_define_rule, []) + | _ => ((num, [num]), phi, role, "", [])) in [(name, role', phi, rule, map (rpair []) deps)] end) 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 --| $$ "," -- (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, []) - | SOME Define_Source => (((num, []), phi), Lemma, zipperposition_define_rule, []) + | SOME (Introduced_Source rule) => (((num, []), phi), Definition, rule, []) + | SOME Define_Source => (((num, []), phi), Definition, zipperposition_define_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 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 full local_name problem = (* Satallax is handled separately, in "atp_satallax.ML". *) if local_name = spassN then parse_spass_line 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 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 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 full local_prover problem tstp = (case core_of_agsyhol_proof tstp of SOME facts => facts |> map (core_inference agsyhol_core_rule) | NONE => 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/ATP/atp_proof_reconstruct.ML b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML --- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML +++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML @@ -1,800 +1,800 @@ (* Title: HOL/Tools/ATP/atp_proof_reconstruct.ML Author: Lawrence C. Paulson, Cambridge University Computer Laboratory Author: Claire Quigley, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Basic proof reconstruction from ATP proofs. *) signature ATP_PROOF_RECONSTRUCT = sig type 'a atp_type = 'a ATP_Problem.atp_type type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula type stature = ATP_Problem_Generate.stature type atp_step_name = ATP_Proof.atp_step_name type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step type 'a atp_proof = 'a ATP_Proof.atp_proof val metisN : string val full_typesN : string val partial_typesN : string val no_typesN : string val really_full_type_enc : string val full_type_enc : string val partial_type_enc : string val no_type_enc : string val full_type_encs : string list val partial_type_encs : string list val default_metis_lam_trans : string val leo2_extcnf_equal_neg_rule : string val satallax_tab_rule_prefix : string val forall_of : term -> term -> term val exists_of : term -> term -> term val simplify_bool : term -> term val var_name_of_typ : typ -> string val rename_bound_vars : term -> term val type_enc_aliases : (string * string list) list val unalias_type_enc : string -> string list val term_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> bool -> int Symtab.table -> typ option -> (string, string atp_type) atp_term -> term val prop_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> bool -> int Symtab.table -> (string, string, (string, string atp_type) atp_term, string) atp_formula -> term val used_facts_in_atp_proof : Proof.context -> (string * stature) list -> string atp_proof -> (string * stature) list val used_facts_in_unsound_atp_proof : Proof.context -> (string * stature) list -> 'a atp_proof -> string list option val atp_proof_prefers_lifting : string atp_proof -> bool val is_typed_helper_used_in_atp_proof : string atp_proof -> bool val replace_dependencies_in_line : atp_step_name * atp_step_name list -> ('a, 'b) atp_step -> ('a, 'b) atp_step val termify_atp_proof : Proof.context -> string -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list -> int Symtab.table -> string atp_proof -> (term, string) atp_step list val repair_waldmeister_endgame : (term, 'a) atp_step list -> (term, 'a) atp_step list val infer_formulas_types : Proof.context -> term list -> term list val introduce_spass_skolems : (term, string) atp_step list -> (term, string) atp_step list val factify_atp_proof : (string * 'a) list -> term list -> term -> (term, string) atp_step list -> (term, string) atp_step list end; structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT = struct open ATP_Util open ATP_Problem open ATP_Proof open ATP_Problem_Generate val metisN = "metis" val full_typesN = "full_types" val partial_typesN = "partial_types" val no_typesN = "no_types" val really_full_type_enc = "mono_tags" val full_type_enc = "poly_guards_query" val partial_type_enc = "poly_args" val no_type_enc = "erased" val full_type_encs = [full_type_enc, really_full_type_enc] val partial_type_encs = partial_type_enc :: full_type_encs val type_enc_aliases = [(full_typesN, full_type_encs), (partial_typesN, partial_type_encs), (no_typesN, [no_type_enc])] fun unalias_type_enc s = AList.lookup (op =) type_enc_aliases s |> the_default [s] val default_metis_lam_trans = combsN val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg" val satallax_tab_rule_prefix = "tab_" fun term_name' (Var ((s, _), _)) = perhaps (try Name.dest_skolem) s | term_name' _ = "" fun lambda' v = Term.lambda_name (term_name' v, v) fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda' v t fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda' v t fun make_tfree ctxt w = let val ww = "'" ^ w in TFree (ww, the_default \<^sort>\type\ (Variable.def_sort ctxt (ww, ~1))) end fun simplify_bool ((all as Const (\<^const_name>\All\, _)) $ Abs (s, T, t)) = let val t' = simplify_bool t in if loose_bvar1 (t', 0) then all $ Abs (s, T, t') else t' end | simplify_bool (Const (\<^const_name>\Not\, _) $ t) = s_not (simplify_bool t) | simplify_bool (Const (\<^const_name>\conj\, _) $ t $ u) = s_conj (simplify_bool t, simplify_bool u) | simplify_bool (Const (\<^const_name>\disj\, _) $ t $ u) = s_disj (simplify_bool t, simplify_bool u) | simplify_bool (Const (\<^const_name>\implies\, _) $ t $ u) = s_imp (simplify_bool t, simplify_bool u) | simplify_bool ((t as Const (\<^const_name>\HOL.eq\, _)) $ u $ v) = (case (u, v) of (Const (\<^const_name>\True\, _), _) => v | (u, Const (\<^const_name>\True\, _)) => u | (Const (\<^const_name>\False\, _), v) => s_not v | (u, Const (\<^const_name>\False\, _)) => s_not u | _ => if u aconv v then \<^Const>\True\ else t $ simplify_bool u $ simplify_bool v) | simplify_bool (t $ u) = simplify_bool t $ simplify_bool u | simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t) | simplify_bool t = t fun single_letter upper s = let val s' = if String.isPrefix "o" s orelse String.isPrefix "O" s then "z" else s in String.extract (Name.desymbolize (SOME upper) (Long_Name.base_name s'), 0, SOME 1) end fun var_name_of_typ (Type (\<^type_name>\fun\, [_, T])) = if body_type T = HOLogic.boolT then "p" else "f" | var_name_of_typ (Type (\<^type_name>\set\, [T])) = let fun default () = single_letter true (var_name_of_typ T) in (case T of Type (s, [T1, T2]) => if String.isSuffix "prod" s andalso T1 = T2 then "r" else default () | _ => default ()) end | var_name_of_typ (Type (s, Ts)) = if String.isSuffix "list" s then var_name_of_typ (the_single Ts) ^ "s" else single_letter false (Long_Name.base_name s) | var_name_of_typ (TFree (s, _)) = single_letter false (perhaps (try (unprefix "'")) s) | var_name_of_typ (TVar ((s, _), T)) = var_name_of_typ (TFree (s, T)) fun rename_bound_vars (t $ u) = rename_bound_vars t $ rename_bound_vars u | rename_bound_vars (Abs (_, T, t)) = Abs (var_name_of_typ T, T, rename_bound_vars t) | rename_bound_vars t = t exception ATP_CLASS of string list exception ATP_TYPE of string atp_type list exception ATP_TERM of (string, string atp_type) atp_term list exception ATP_FORMULA of (string, string, (string, string atp_type) atp_term, string) atp_formula list exception SAME of unit fun class_of_atp_class cls = (case unprefix_and_unascii class_prefix cls of SOME s => s | NONE => raise ATP_CLASS [cls]) fun atp_type_of_atp_term (ATerm ((s, _), us)) = let val tys = map atp_type_of_atp_term us in if s = tptp_fun_type then (case tys of [ty1, ty2] => AFun (ty1, ty2) | _ => raise ATP_TYPE tys) else AType ((s, []), tys) end (* Type variables are given the basic sort "HOL.type". Some will later be constrained by information from type literals, or by type inference. *) fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) = let val Ts = map (typ_of_atp_type ctxt) tys in (case unprefix_and_unascii native_type_prefix a of SOME b => typ_of_atp_type ctxt (atp_type_of_atp_term (unmangled_type b)) | NONE => (case unprefix_and_unascii type_const_prefix a of SOME b => Type (invert_const b, Ts) | NONE => if not (null tys) then raise ATP_TYPE [ty] (* only "tconst"s have type arguments *) else (case unprefix_and_unascii tfree_prefix a of SOME b => make_tfree ctxt b | NONE => (* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018". Sometimes variables from the ATP are indistinguishable from Isabelle variables, which forces us to use a type parameter in all cases. *) Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a, (if null clss then \<^sort>\type\ else map class_of_atp_class clss))))) end | typ_of_atp_type ctxt (AFun (ty1, ty2)) = typ_of_atp_type ctxt ty1 --> typ_of_atp_type ctxt ty2 fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term (* Type class literal applied to a type. Returns triple of polarity, class, type. *) fun type_constraint_of_term ctxt (u as ATerm ((a, _), us)) = (case (unprefix_and_unascii class_prefix a, map (typ_of_atp_term ctxt) us) of (SOME b, [T]) => (b, T) | _ => raise ATP_TERM [u]) (* Accumulate type constraints in a formula: negative type literals. *) fun add_var (key, z) = Vartab.map_default (key, []) (cons z) fun add_type_constraint false (cl, TFree (a ,_)) = add_var ((a, ~1), cl) | add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl) | add_type_constraint _ _ = I fun repair_var_name s = (case unprefix_and_unascii schematic_var_prefix s of SOME s' => s' | NONE => s) (* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem pseudoconstants, this information is encoded in the constant name. *) fun robust_const_num_type_args thy s = if String.isPrefix skolem_const_prefix s then s |> Long_Name.explode |> List.last |> Int.fromString |> the else if String.isPrefix lam_lifted_prefix s then if String.isPrefix lam_lifted_poly_prefix s then 2 else 0 else (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length fun slack_fastype_of t = fastype_of t handle TERM _ => Type_Infer.anyT \<^sort>\type\ val spass_skolem_prefix = "sk" (* "skc" or "skf" *) val vampire_skolem_prefix = "sK" fun var_index_of_textual textual = if textual then 0 else 1 fun quantify_over_var textual quant_of var_s var_T t = let val vars = ((var_s, var_index_of_textual textual), var_T) :: filter (fn ((s, _), _) => s = var_s) (Term.add_vars t []) val normTs = vars |> AList.group (op =) |> map (apsnd hd) fun norm_var_types (Var (x, T)) = Var (x, the_default T (AList.lookup (op =) normTs x)) | norm_var_types t = t in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end (* This assumes that distinct names are mapped to distinct names by "Variable.variant_frees". This does not hold in general but should hold for ATP-generated Skolem function names, since these end with a digit and "variant_frees" appends letters. *) fun fresh_up ctxt s = fst (hd (Variable.variant_frees ctxt [] [(s, ())])) (* Higher-order translation. Variables are typed (although we don't use that information). Lambdas are typed. The code is similar to "term_of_atp_fo". *) fun term_of_atp_ho ctxt sym_tab = let val thy = Proof_Context.theory_of ctxt val var_index = var_index_of_textual true fun do_term opt_T u = (case u of AAbs (((var, ty), term), []) => let val typ = typ_of_atp_type ctxt ty val var_name = repair_var_name var val tm = do_term NONE term in quantify_over_var true lambda' var_name typ tm end | ATerm ((s, tys), us) => if s = "" (* special marker generated on parse error *) then error "Isar proof reconstruction failed because the ATP proof contains unparsable \ \material" else if s = tptp_equal then list_comb (Const (\<^const_name>\HOL.eq\, Type_Infer.anyT \<^sort>\type\), map (do_term NONE) us) else if s = tptp_not_equal andalso length us = 2 then \<^const>\HOL.Not\ $ list_comb (Const (\<^const_name>\HOL.eq\, Type_Infer.anyT \<^sort>\type\), map (do_term NONE) us) else if not (null us) then let val args = map (do_term NONE) us val opt_T' = SOME (map slack_fastype_of args ---> the_default dummyT opt_T) val func = do_term opt_T' (ATerm ((s, tys), [])) in foldl1 (op $) (func :: args) end else if s = tptp_or then HOLogic.disj else if s = tptp_and then HOLogic.conj else if s = tptp_implies then HOLogic.imp else if s = tptp_iff orelse s = tptp_equal then HOLogic.eq_const dummyT else if s = tptp_not_iff orelse s = tptp_not_equal then \<^term>\\x y. x \ y\ else if s = tptp_if then \<^term>\\P Q. Q \ P\ else if s = tptp_not_and then \<^term>\\P Q. \ (P \ Q)\ else if s = tptp_not_or then \<^term>\\P Q. \ (P \ Q)\ else if s = tptp_not then HOLogic.Not else if s = tptp_ho_forall then HOLogic.all_const dummyT else if s = tptp_ho_exists then HOLogic.exists_const dummyT else if s = tptp_hilbert_choice then HOLogic.choice_const dummyT else if s = tptp_hilbert_the then \<^term>\The\ else (case unprefix_and_unascii const_prefix s of SOME s' => let val ((s', _), mangled_us) = s' |> unmangled_const |>> `invert_const val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) val (type_us, term_us) = chop num_ty_args us |>> append mangled_us val term_ts = map (do_term NONE) term_us val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us val T = (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then try (Sign.const_instance thy) (s', Ts) else NONE) |> (fn SOME T => T | NONE => map slack_fastype_of term_ts ---> the_default (Type_Infer.anyT \<^sort>\type\) opt_T) val t = Const (unproxify_const s', T) in list_comb (t, term_ts) end | NONE => (* a free or schematic variable *) let val ts = map (do_term NONE) us val T = (case tys of [ty] => typ_of_atp_type ctxt ty | _ => map slack_fastype_of ts ---> (case opt_T of SOME T => T | NONE => Type_Infer.anyT \<^sort>\type\)) val t = (case unprefix_and_unascii fixed_var_prefix s of SOME s => Free (s, T) | NONE => if not (is_tptp_variable s) then Free (fresh_up ctxt s, T) else Var ((repair_var_name s, var_index), T)) in list_comb (t, ts) end)) in do_term end (* First-order translation. No types are known for variables. "Type_Infer.anyT @{sort type}" should allow them to be inferred. *) fun term_of_atp_fo ctxt textual sym_tab = let val thy = Proof_Context.theory_of ctxt (* For Metis, we use 1 rather than 0 because variable references in clauses may otherwise conflict with variable constraints in the goal. At least, type inference often fails otherwise. See also "axiom_inference" in "Metis_Reconstruct". *) val var_index = var_index_of_textual textual fun do_term extra_ts opt_T u = (case u of ATerm ((s, tys), us) => if s = "" (* special marker generated on parse error *) then error "Isar proof reconstruction failed because the ATP proof contains unparsable \ \material" else if String.isPrefix native_type_prefix s then \<^Const>\True\ (* ignore TPTP type information (needed?) *) else if s = tptp_equal then list_comb (Const (\<^const_name>\HOL.eq\, Type_Infer.anyT \<^sort>\type\), map (do_term [] NONE) us) else (case unprefix_and_unascii const_prefix s of SOME s' => let val ((s', s''), mangled_us) = s' |> unmangled_const |>> `invert_const in if s' = type_tag_name then (case mangled_us @ us of [typ_u, term_u] => do_term extra_ts (SOME (typ_of_atp_term ctxt typ_u)) term_u | _ => raise ATP_TERM us) else if s' = predicator_name then do_term [] (SOME \<^typ>\bool\) (hd us) else if s' = app_op_name then let val extra_t = do_term [] NONE (List.last us) in do_term (extra_t :: extra_ts) (case opt_T of SOME T => SOME (slack_fastype_of extra_t --> T) | NONE => NONE) (nth us (length us - 2)) end else if s' = type_guard_name then \<^Const>\True\ (* ignore type predicates *) else let val new_skolem = String.isPrefix new_skolem_const_prefix s'' val num_ty_args = length us - the_default 0 (Symtab.lookup sym_tab s) val (type_us, term_us) = chop num_ty_args us |>> append mangled_us val term_ts = map (do_term [] NONE) term_us val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us val T = (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then if new_skolem then SOME (Type_Infer.paramify_vars (tl Ts ---> hd Ts)) else if textual then try (Sign.const_instance thy) (s', Ts) else NONE else NONE) |> (fn SOME T => T | NONE => map slack_fastype_of term_ts ---> the_default (Type_Infer.anyT \<^sort>\type\) opt_T) val t = if new_skolem then Var ((new_skolem_var_name_of_const s'', var_index), T) else Const (unproxify_const s', T) in list_comb (t, term_ts @ extra_ts) end end | NONE => (* a free or schematic variable *) let val term_ts = map (do_term [] NONE) us (* SPASS (3.8ds) and Vampire (2.6) pass arguments to Skolem functions in reverse order, which is incompatible with "metis"'s new skolemizer. *) |> exists (fn pre => String.isPrefix pre s) [spass_skolem_prefix, vampire_skolem_prefix] ? rev val ts = term_ts @ extra_ts val T = (case tys of [ty] => typ_of_atp_type ctxt ty | _ => (case opt_T of SOME T => map slack_fastype_of term_ts ---> T | NONE => map slack_fastype_of ts ---> Type_Infer.anyT \<^sort>\type\)) val t = (case unprefix_and_unascii fixed_var_prefix s of SOME s => Free (s, T) | NONE => if textual andalso not (is_tptp_variable s) then Free (s |> textual ? fresh_up ctxt, T) else Var ((repair_var_name s, var_index), T)) in list_comb (t, ts) end)) in do_term [] end fun term_of_atp ctxt (ATP_Problem.THF _) type_enc = if ATP_Problem_Generate.is_type_enc_higher_order type_enc then K (term_of_atp_ho ctxt) else error "Unsupported Isar reconstruction" | term_of_atp ctxt _ type_enc = if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc) then term_of_atp_fo ctxt else error "Unsupported Isar reconstruction" fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) = if String.isPrefix class_prefix s then add_type_constraint pos (type_constraint_of_term ctxt u) #> pair \<^Const>\True\ else pair (term_of_atp ctxt format type_enc textual sym_tab (SOME \<^typ>\bool\) u) (* Update schematic type variables with detected sort constraints. It's not totally clear whether this code is necessary. *) fun repair_tvar_sorts (t, tvar_tab) = let fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) | do_type (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) | do_type (TFree z) = TFree z fun do_term (Const (a, T)) = Const (a, do_type T) | do_term (Free (a, T)) = Free (a, do_type T) | do_term (Var (xi, T)) = Var (xi, do_type T) | do_term (t as Bound _) = t | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) | do_term (t1 $ t2) = do_term t1 $ do_term t2 in t |> not (Vartab.is_empty tvar_tab) ? do_term end (* Interpret an ATP formula as a HOL term, extracting sort constraints as they appear in the formula. *) fun prop_of_atp ctxt format type_enc textual sym_tab phi = let fun do_formula pos phi = (case phi of AQuant (_, [], phi) => do_formula pos phi | AQuant (q, (s, _) :: xs, phi') => do_formula pos (AQuant (q, xs, phi')) (* FIXME: TFF *) #>> quantify_over_var textual (case q of AForall => forall_of | AExists => exists_of) (repair_var_name s) dummyT | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not | AConn (c, [phi1, phi2]) => do_formula (pos |> c = AImplies ? not) phi1 ##>> do_formula pos phi2 #>> (case c of AAnd => s_conj | AOr => s_disj | AImplies => s_imp | AIff => s_iff | ANot => raise Fail "impossible connective") | AAtom tm => term_of_atom ctxt format type_enc textual sym_tab pos tm | _ => raise ATP_FORMULA [phi]) in repair_tvar_sorts (do_formula true phi Vartab.empty) end val unprefix_fact_number = space_implode "_" o tl o space_explode "_" fun resolve_fact facts s = (case try (unprefix fact_prefix) s of SOME s' => let val s' = s' |> unprefix_fact_number |> unascii_of in AList.lookup (op =) facts s' |> Option.map (pair s') end | NONE => NONE) fun resolve_conjecture s = (case try (unprefix conjecture_prefix) s of SOME s' => Int.fromString s' | NONE => NONE) fun resolve_facts facts = map_filter (resolve_fact facts) val resolve_conjectures = map_filter resolve_conjecture fun is_axiom_used_in_proof pred = exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false) fun add_fact ctxt facts ((num, ss), _, _, rule, deps) = (if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule] rule orelse String.isPrefix satallax_tab_rule_prefix rule then insert (op =) (short_thm_name ctxt ext, (Global, General)) else I) #> (if null deps then union (op =) (resolve_facts facts (num :: ss)) else I) fun used_facts_in_atp_proof ctxt facts atp_proof = if null atp_proof then facts else fold (add_fact ctxt facts) atp_proof [] fun used_facts_in_unsound_atp_proof _ _ [] = NONE | used_facts_in_unsound_atp_proof ctxt facts atp_proof = let val used_facts = used_facts_in_atp_proof ctxt facts atp_proof val all_global_facts = forall (fn (_, (sc, _)) => sc = Global) used_facts val axiom_used = is_axiom_used_in_proof (is_some o resolve_conjecture) atp_proof in if all_global_facts andalso not axiom_used then SOME (map fst used_facts) else NONE end val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix (* overapproximation (good enough) *) fun is_lam_lifted s = String.isPrefix fact_prefix s andalso String.isSubstring ascii_of_lam_fact_prefix s val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix) fun atp_proof_prefers_lifting atp_proof = (is_axiom_used_in_proof is_combinator_def atp_proof, is_axiom_used_in_proof is_lam_lifted atp_proof) = (false, true) val is_typed_helper_name = String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix fun is_typed_helper_used_in_atp_proof atp_proof = is_axiom_used_in_proof is_typed_helper_name atp_proof fun replace_one_dependency (old, new) dep = if is_same_atp_step dep old then new else [dep] fun replace_dependencies_in_line old_new (name, role, t, rule, deps) = (name, role, t, rule, fold (union (op =) o replace_one_dependency old_new) deps []) fun repair_name "$true" = "c_True" | repair_name "$false" = "c_False" | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *) | repair_name s = if is_tptp_equal s orelse (* seen in Vampire proofs *) (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then tptp_equal else s fun set_var_index j = map_aterms (fn Var ((s, 0), T) => Var ((s, j), T) | t => t) fun infer_formulas_types ctxt = map_index (uncurry (fn j => set_var_index j #> Type.constraint HOLogic.boolT)) #> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) val combinator_table = [(\<^const_name>\Meson.COMBI\, @{thm Meson.COMBI_def [abs_def]}), (\<^const_name>\Meson.COMBK\, @{thm Meson.COMBK_def [abs_def]}), (\<^const_name>\Meson.COMBB\, @{thm Meson.COMBB_def [abs_def]}), (\<^const_name>\Meson.COMBC\, @{thm Meson.COMBC_def [abs_def]}), (\<^const_name>\Meson.COMBS\, @{thm Meson.COMBS_def [abs_def]})] fun uncombine_term thy = let fun uncomb (t1 $ t2) = betapply (uncomb t1, uncomb t2) | uncomb (Abs (s, T, t)) = Abs (s, T, uncomb t) | uncomb (t as Const (x as (s, _))) = (case AList.lookup (op =) combinator_table s of SOME thm => thm |> Thm.prop_of |> specialize_type thy x |> Logic.dest_equals |> snd | NONE => t) | uncomb t = t in uncomb end fun unlift_aterm lifted (t as Const (s, _)) = if String.isPrefix lam_lifted_prefix s then (* FIXME: do something about the types *) (case AList.lookup (op =) lifted s of SOME t' => unlift_term lifted t' | NONE => t) else t | unlift_aterm _ t = t and unlift_term lifted = map_aterms (unlift_aterm lifted) fun termify_line _ _ _ _ _ (_, Type_Role, _, _, _) = NONE | termify_line ctxt format type_enc lifted sym_tab (name, role, u, rule, deps) = let val thy = Proof_Context.theory_of ctxt val t = u |> prop_of_atp ctxt format type_enc true sym_tab |> unlift_term lifted |> uncombine_term thy |> simplify_bool in SOME (name, role, t, rule, deps) end val waldmeister_conjecture_num = "1.0.0.0" fun repair_waldmeister_endgame proof = let fun repair_tail (name, _, \<^Const>\Trueprop for t\, rule, deps) = (name, Negated_Conjecture, \<^Const>\Trueprop for \s_not t\\, rule, deps) fun repair_body [] = [] | repair_body ((line as ((num, _), _, _, _, _)) :: lines) = if num = waldmeister_conjecture_num then map repair_tail (line :: lines) else line :: repair_body lines in repair_body proof end fun map_proof_terms f (lines : ('a * 'b * 'c * 'd * 'e) list) = map2 (fn c => fn (a, b, _, d, e) => (a, b, c, d, e)) (f (map #3 lines)) lines fun termify_atp_proof ctxt local_prover format type_enc pool lifted sym_tab = nasty_atp_proof pool #> map_term_names_in_atp_proof repair_name #> map_filter (termify_line ctxt format type_enc lifted sym_tab) #> map_proof_terms (infer_formulas_types ctxt #> map HOLogic.mk_Trueprop) #> local_prover = waldmeisterN ? repair_waldmeister_endgame fun unskolemize_term skos = let val is_skolem_name = member (op =) skos fun find_argless_skolem (Free _ $ Var _) = NONE | find_argless_skolem (Free (x as (s, _))) = if is_skolem_name s then SOME x else NONE | find_argless_skolem (t $ u) = (case find_argless_skolem t of NONE => find_argless_skolem u | sk => sk) | find_argless_skolem (Abs (_, _, t)) = find_argless_skolem t | find_argless_skolem _ = NONE fun find_skolem_arg (Free (s, _) $ Var v) = if is_skolem_name s then SOME v else NONE | find_skolem_arg (t $ u) = (case find_skolem_arg t of NONE => find_skolem_arg u | v => v) | find_skolem_arg (Abs (_, _, t)) = find_skolem_arg t | find_skolem_arg _ = NONE fun kill_skolem_arg (t as Free (s, T) $ Var _) = if is_skolem_name s then Free (s, range_type T) else t | kill_skolem_arg (t $ u) = kill_skolem_arg t $ kill_skolem_arg u | kill_skolem_arg (Abs (s, T, t)) = Abs (s, T, kill_skolem_arg t) | kill_skolem_arg t = t fun find_var (Var v) = SOME v | find_var (t $ u) = (case find_var t of NONE => find_var u | v => v) | find_var (Abs (_, _, t)) = find_var t | find_var _ = NONE val safe_abstract_over = abstract_over o apsnd (incr_boundvars 1) fun unskolem t = (case find_argless_skolem t of SOME (x as (s, T)) => HOLogic.exists_const T $ Abs (s, T, unskolem (safe_abstract_over (Free x, t))) | NONE => (case find_skolem_arg t of SOME (v as ((s, _), T)) => let val (haves, have_nots) = HOLogic.disjuncts t |> List.partition (exists_subterm (curry (op =) (Var v))) |> apply2 (fn lits => fold (curry s_disj) lits \<^term>\False\) in s_disj (HOLogic.all_const T $ Abs (s, T, unskolem (safe_abstract_over (Var v, kill_skolem_arg haves))), unskolem have_nots) end | NONE => (case find_var t of SOME (v as ((s, _), T)) => HOLogic.all_const T $ Abs (s, T, unskolem (safe_abstract_over (Var v, t))) | NONE => t))) in HOLogic.mk_Trueprop o unskolem o HOLogic.dest_Trueprop end fun rename_skolem_args t = let fun add_skolem_args (Abs (_, _, t)) = add_skolem_args t | add_skolem_args t = (case strip_comb t of (Free (s, _), args as _ :: _) => if String.isPrefix spass_skolem_prefix s then insert (op =) (s, take_prefix is_Var args) else fold add_skolem_args args | (u, args as _ :: _) => fold add_skolem_args (u :: args) | _ => I) fun subst_of_skolem (sk, args) = map_index (fn (j, Var (z, T)) => (z, Var ((sk ^ "_" ^ string_of_int j, 0), T))) args val subst = maps subst_of_skolem (add_skolem_args t []) in subst_vars ([], subst) t end fun introduce_spass_skolems proof = let fun add_skolem (Free (s, _)) = String.isPrefix spass_skolem_prefix s ? insert (op =) s | add_skolem _ = I (* union-find would be faster *) fun add_cycle [] = I | add_cycle ss = fold (fn s => Graph.default_node (s, ())) ss #> fold Graph.add_edge (ss ~~ tl ss @ [hd ss]) val (input_steps, other_steps) = List.partition (null o #5) proof (* The names of the formulas are added to the Skolem constants, to ensure that formulas giving rise to several clauses are skolemized together. *) val skoXss = map (fn ((_, ss), _, t, _, _) => Term.fold_aterms add_skolem t ss) input_steps val groups0 = Graph.strong_conn (fold add_cycle skoXss Graph.empty) val groups = filter (exists (String.isPrefix spass_skolem_prefix)) groups0 val skoXss_input_steps = skoXss ~~ input_steps fun step_name_of_group skoXs = (implode skoXs, []) fun in_group group = member (op =) group o hd fun group_of skoX = find_first (fn group => in_group group skoX) groups fun new_steps (skoXss_steps : (string list * (term, 'a) atp_step) list) group = let val name = step_name_of_group group val name0 = name |>> prefix "0" val t = (case map (snd #> #3) skoXss_steps of [t] => t | ts => ts |> map (HOLogic.dest_Trueprop #> rename_skolem_args) |> Library.foldr1 s_conj |> HOLogic.mk_Trueprop) val skos = fold (union (op =) o filter (String.isPrefix spass_skolem_prefix) o fst) skoXss_steps [] val deps = map (snd #> #1) skoXss_steps in [(name0, Unknown, unskolemize_term skos t, spass_pre_skolemize_rule, deps), (name, Unknown, t, spass_skolemize_rule, [name0])] end val sko_steps = maps (fn group => new_steps (filter (in_group group o fst) skoXss_input_steps) group) groups val old_news = map_filter (fn (skoXs, (name, _, _, _, _)) => Option.map (pair name o single o step_name_of_group) (group_of skoXs)) skoXss_input_steps val repair_deps = fold replace_dependencies_in_line old_news in input_steps @ sko_steps @ map repair_deps other_steps end fun factify_atp_proof facts hyp_ts concl_t atp_proof = let fun factify_step ((num, ss), role, t, rule, deps) = let val (ss', role', t') = (case resolve_conjectures ss of [j] => if j = length hyp_ts then ([], Conjecture, concl_t) else ([], Hypothesis, close_form (nth hyp_ts j)) | _ => (case resolve_facts facts (num :: ss) of - [] => (ss, if role = Lemma then Lemma else Plain, t) + [] => (ss, if member (op =) [Definition, Lemma] role then role else Plain, t) | facts => (map fst facts, Axiom, t))) in ((num, ss'), role', t', rule, deps) end val atp_proof = map factify_step atp_proof val names = map #1 atp_proof fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num)) fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps) in map repair_deps atp_proof end end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML @@ -1,523 +1,526 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_isar.ML Author: Jasmin Blanchette, TU Muenchen Author: Steffen Juilf Smolka, TU Muenchen Isar proof reconstruction from ATP proofs. *) signature SLEDGEHAMMER_ISAR = sig type atp_step_name = ATP_Proof.atp_step_name type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step type 'a atp_proof = 'a ATP_Proof.atp_proof type stature = ATP_Problem_Generate.stature type one_line_params = Sledgehammer_Proof_Methods.one_line_params val trace : bool Config.T type isar_params = bool * (string option * string option) * Time.time * real option * bool * bool * (term, string) atp_step list * thm val proof_text : Proof.context -> bool -> bool option -> bool -> (unit -> isar_params) -> int -> one_line_params -> string end; structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR = 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_Proof open Sledgehammer_Isar_Preplay open Sledgehammer_Isar_Compress open Sledgehammer_Isar_Minimize structure String_Redirect = ATP_Proof_Redirect( type key = atp_step_name val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') val string_of = fst) open String_Redirect val trace = Attrib.setup_config_bool \<^binding>\sledgehammer_isar_trace\ (K false) val e_definition_rule = "definition" val e_skolemize_rule = "skolemize" val leo2_extcnf_forall_neg_rule = "extcnf_forall_neg" val satallax_skolemize_rule = "tab_ex" val vampire_skolemisation_rule = "skolemisation" val veriT_la_generic_rule = "la_generic" val veriT_simp_arith_rule = "simp_arith" val veriT_skolemize_rules = Verit_Proof.skolemization_steps val z3_skolemize_rule = Z3_Proof.string_of_rule Z3_Proof.Skolemize val z3_th_lemma_rule_prefix = Z3_Proof.string_of_rule (Z3_Proof.Th_Lemma "") val zipperposition_cnf_rule = "cnf" val skolemize_rules = [e_definition_rule, e_skolemize_rule, leo2_extcnf_forall_neg_rule, satallax_skolemize_rule, spass_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule, zipperposition_cnf_rule] @ veriT_skolemize_rules fun is_ext_rule rule = (rule = leo2_extcnf_equal_neg_rule) val is_maybe_ext_rule = is_ext_rule orf String.isPrefix satallax_tab_rule_prefix val is_skolemize_rule = member (op =) skolemize_rules fun is_arith_rule rule = String.isPrefix z3_th_lemma_rule_prefix rule orelse rule = veriT_simp_arith_rule orelse rule = veriT_la_generic_rule fun raw_label_of_num num = (num, 0) fun label_of_clause [(num, _)] = raw_label_of_num num | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0) fun add_global_fact ss = apsnd (union (op =) ss) fun add_fact_of_dependency [(_, ss as _ :: _)] = add_global_fact ss | add_fact_of_dependency names = apfst (insert (op =) (label_of_clause names)) fun add_line_pass1 (line as (name, role, t, rule, [])) lines = (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or definitions. *) if role = Conjecture orelse role = Negated_Conjecture then line :: lines else if t aconv \<^prop>\True\ then map (replace_dependencies_in_line (name, [])) lines - else if role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then + else if role = Definition orelse role = Lemma orelse role = Hypothesis + orelse is_arith_rule rule then line :: lines else if role = Axiom then lines (* axioms (facts) need no proof lines *) else map (replace_dependencies_in_line (name, [])) lines | add_line_pass1 line lines = line :: lines fun add_lines_pass2 res [] = rev res | add_lines_pass2 res ((line as (name, role, t, rule, deps)) :: lines) = let fun normalize role = role = Conjecture ? (HOLogic.dest_Trueprop #> s_not #> HOLogic.mk_Trueprop) val norm_t = normalize role t val is_duplicate = exists (fn (prev_name, prev_role, prev_t, _, _) => (prev_role = Hypothesis andalso prev_t aconv t) orelse (member (op =) deps prev_name andalso Term.aconv_untyped (normalize prev_role prev_t, norm_t))) res fun looks_boring () = t aconv \<^prop>\False\ orelse length deps < 2 fun is_skolemizing_line (_, _, _, rule', deps') = is_skolemize_rule rule' andalso member (op =) deps' name fun is_before_skolemize_rule () = exists is_skolemizing_line lines in if is_duplicate orelse (role = Plain andalso not (is_skolemize_rule rule) andalso not (is_ext_rule rule) andalso not (is_arith_rule rule) andalso not (null lines) andalso looks_boring () andalso not (is_before_skolemize_rule ())) then add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines) else add_lines_pass2 (line :: res) lines end type isar_params = bool * (string option * string option) * Time.time * real option * bool * bool * (term, string) atp_step list * thm val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method, SATx_Method] val basic_simp_based_methods = [Auto_Method, Simp_Method, Fastforce_Method, Force_Method] val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method] val arith_methods = basic_arith_methods @ basic_simp_based_methods @ basic_systematic_methods val systematic_methods = basic_systematic_methods @ basic_arith_methods @ basic_simp_based_methods @ [Metis_Method (SOME full_typesN, NONE), Metis_Method (SOME no_typesN, NONE)] val rewrite_methods = basic_simp_based_methods @ basic_systematic_methods @ basic_arith_methods val skolem_methods = Moura_Method :: systematic_methods fun isar_proof_text ctxt debug num_chained isar_proofs smt_proofs isar_params (one_line_params as ((used_facts, (_, one_line_play)), banner, subgoal, subgoal_count)) = let val _ = if debug then writeln "Constructing Isar proof..." else () fun generate_proof_text () = let val (verbose, alt_metis_args, preplay_timeout, compress, try0, minimize, atp_proof0, goal) = isar_params () in if null atp_proof0 then one_line_proof_text ctxt 0 one_line_params else let val systematic_methods' = insert (op =) (Metis_Method alt_metis_args) systematic_methods fun massage_methods (meths as meth :: _) = if not try0 then [meth] else if smt_proofs then SMT_Method SMT_Z3 :: meths else meths val (params, _, concl_t) = strip_subgoal goal subgoal ctxt val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params val ctxt = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd val do_preplay = preplay_timeout <> Time.zeroTime val compress = (case compress of NONE => if isar_proofs = NONE andalso do_preplay then 1000.0 else 10.0 | SOME n => n) fun is_fixed ctxt = Variable.is_declared ctxt orf Name.is_skolem fun skolems_of ctxt t = Term.add_frees t [] |> filter_out (is_fixed ctxt o fst) |> rev fun get_role keep_role ((num, _), role, t, rule, _) = if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE val trace = Config.get ctxt trace val string_of_atp_steps = let val to_string = ATP_Proof.string_of_atp_step (Syntax.string_of_term ctxt) I in enclose "[\n" "\n]" o cat_lines o map (enclose " " "," o to_string) end val atp_proof = atp_proof0 |> trace ? tap (tracing o prefix "atp_proof0 = " o string_of_atp_steps) |> distinct (op =) (* Zipperposition generates duplicate lines *) |> (fn lines => fold_rev add_line_pass1 lines []) |> add_lines_pass2 [] |> trace ? tap (tracing o prefix "atp_proof = " o string_of_atp_steps) val conjs = map_filter (fn (name, role, _, _, _) => if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE) atp_proof val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof fun add_lemma ((label, goal), rule) ctxt = let val (skos, proof_methods) = (if is_skolemize_rule rule then (skolems_of ctxt goal, skolem_methods) else if is_arith_rule rule then ([], arith_methods) else ([], rewrite_methods)) ||> massage_methods val prove = Prove { qualifiers = [], obtains = skos, label = label, goal = goal, subproofs = [], facts = ([], []), proof_methods = proof_methods, comment = ""} in (prove, ctxt |> not (null skos) ? (Variable.add_fixes (map fst skos) #> snd)) end val (lems, _) = - fold_map add_lemma (map_filter (get_role (curry (op =) Lemma)) atp_proof) ctxt + fold_map add_lemma (map_filter (get_role (member (op =) [Definition, Lemma])) + atp_proof) ctxt val bot = #1 (List.last atp_proof) val refute_graph = atp_proof |> map (fn (name, _, _, _, from) => (from, name)) |> make_refute_graph bot |> fold (Atom_Graph.default_node o rpair ()) conjs val axioms = axioms_of_refute_graph refute_graph conjs val tainted = tainted_atoms_of_refute_graph refute_graph conjs val is_clause_tainted = exists (member (op =) tainted) val steps = Symtab.empty |> fold (fn (name as (s, _), role, t, rule, _) => Symtab.update_new (s, (rule, t |> (if is_clause_tainted [name] then HOLogic.dest_Trueprop #> role <> Conjecture ? s_not #> fold exists_of (map Var (Term.add_vars t [])) #> HOLogic.mk_Trueprop else I)))) atp_proof fun is_referenced_in_step _ (Let _) = false | is_referenced_in_step l (Prove {subproofs, facts = (ls, _), ...}) = member (op =) ls l orelse exists (is_referenced_in_proof l) subproofs and is_referenced_in_proof l (Proof {steps, ...}) = exists (is_referenced_in_step l) steps fun insert_lemma_in_step lem (step as Prove {qualifiers, obtains, label, goal, subproofs, facts = (ls, gs), proof_methods, comment}) = let val l' = the (label_of_isar_step lem) in if member (op =) ls l' then [lem, step] else let val refs = map (is_referenced_in_proof l') subproofs in if length (filter I refs) = 1 then [Prove { qualifiers = qualifiers, obtains = obtains, label = label, goal = goal, subproofs = map2 (fn false => I | true => insert_lemma_in_proof lem) refs subproofs, facts = (ls, gs), proof_methods = proof_methods, comment = comment}] else [lem, step] end end and insert_lemma_in_steps lem [] = [lem] | insert_lemma_in_steps lem (step :: steps) = if is_referenced_in_step (the (label_of_isar_step lem)) step then insert_lemma_in_step lem step @ steps else step :: insert_lemma_in_steps lem steps and insert_lemma_in_proof lem (proof as Proof {steps, ...}) = isar_proof_with_steps proof (insert_lemma_in_steps lem steps) val rule_of_clause_id = fst o the o Symtab.lookup steps o fst val finish_off = close_form #> rename_bound_vars fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> finish_off | prop_of_clause names = let val lits = map (HOLogic.dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names) in (case List.partition (can HOLogic.dest_not) lits of (negs as _ :: _, pos as _ :: _) => s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) | _ => fold (curry s_disj) lits \<^term>\False\) end |> HOLogic.mk_Trueprop |> finish_off fun maybe_show outer c = if outer andalso eq_set (op =) (c, conjs) then [Show] else [] fun isar_steps outer predecessor accum [] = accum |> (if tainted = [] then (* e.g., trivial, empty proof by Z3 *) cons (Prove { qualifiers = if outer then [Show] else [], obtains = [], label = no_label, goal = concl_t, subproofs = [], facts = sort_facts (the_list predecessor, []), proof_methods = massage_methods systematic_methods', comment = ""}) else I) |> rev | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) = let val l = label_of_clause c val t = prop_of_clause c val rule = rule_of_clause_id id val skolem = is_skolemize_rule rule val deps = ([], []) |> fold add_fact_of_dependency gamma |> is_maybe_ext_rule rule ? add_global_fact [short_thm_name ctxt ext] |> sort_facts val meths = (if skolem then skolem_methods else if is_arith_rule rule then arith_methods else systematic_methods') |> massage_methods fun prove subproofs facts = Prove { qualifiers = maybe_show outer c, obtains = [], label = l, goal = t, subproofs = subproofs, facts = facts, proof_methods = meths, comment = ""} fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs in if is_clause_tainted c then (case gamma of [g] => if skolem andalso is_clause_tainted g andalso not (null accum) then let val skos = skolems_of ctxt (prop_of_clause g) val subproof = Proof {fixes = skos, assumptions = [], steps = rev accum} in isar_steps outer (SOME l) [prove [subproof] ([], [])] infs end else steps_of_rest (prove [] deps) | _ => steps_of_rest (prove [] deps)) else steps_of_rest (if skolem then (case skolems_of ctxt t of [] => prove [] deps | skos => Prove { qualifiers = [], obtains = skos, label = l, goal = t, subproofs = [], facts = deps, proof_methods = meths, comment = ""}) else prove [] deps) end | isar_steps outer predecessor accum (Cases cases :: infs) = let fun isar_case (c, subinfs) = isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs val c = succedent_of_cases cases val l = label_of_clause c val t = prop_of_clause c val step = Prove { qualifiers = maybe_show outer c, obtains = [], label = l, goal = t, subproofs = map isar_case (filter_out (null o snd) cases), facts = sort_facts (the_list predecessor, []), proof_methods = massage_methods systematic_methods', comment = ""} in isar_steps outer (SOME l) (step :: accum) infs end and isar_proof outer fixes assumptions lems infs = let val steps = fold_rev insert_lemma_in_steps lems (isar_steps outer NONE [] infs) in Proof {fixes = fixes, assumptions = assumptions, steps = steps} end val canonical_isar_proof = refute_graph |> trace ? tap (tracing o prefix "Refute graph:\n" o string_of_refute_graph) |> redirect_graph axioms tainted bot |> trace ? tap (tracing o prefix "Direct proof:\n" o string_of_direct_proof) |> isar_proof true params assms lems |> postprocess_isar_proof_remove_show_stuttering |> postprocess_isar_proof_remove_unreferenced_steps I |> relabel_isar_proof_canonically val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty val _ = fold_isar_steps (fn meth => K (set_preplay_outcomes_of_isar_step ctxt preplay_timeout preplay_data meth [])) (steps_of_isar_proof canonical_isar_proof) () fun str_of_preplay_outcome outcome = if Lazy.is_finished outcome then string_of_play_outcome (Lazy.force outcome) else "?" fun str_of_meth l meth = string_of_proof_method [] meth ^ " " ^ str_of_preplay_outcome (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth) fun comment_of l = map (str_of_meth l) #> commas fun trace_isar_proof label proof = if trace then tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^ string_of_isar_proof ctxt subgoal subgoal_count (comment_isar_proof comment_of proof) ^ "\n") else () fun comment_of l (meth :: _) = (case (verbose, Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)) of (false, Played _) => "" | (_, outcome) => string_of_play_outcome outcome) val (play_outcome, isar_proof) = canonical_isar_proof |> tap (trace_isar_proof "Original") |> compress_isar_proof ctxt compress preplay_timeout preplay_data |> tap (trace_isar_proof "Compressed") |> postprocess_isar_proof_remove_unreferenced_steps (do_preplay ? keep_fastest_method_of_isar_step (!preplay_data) #> minimize ? minimize_isar_step_dependencies ctxt preplay_data) |> tap (trace_isar_proof "Minimized") - |> `(preplay_outcome_of_isar_proof (!preplay_data)) + |> `(if do_preplay then preplay_outcome_of_isar_proof (!preplay_data) + else K (Play_Timed_Out Time.zeroTime)) ||> (comment_isar_proof comment_of #> chain_isar_proof #> kill_useless_labels_in_isar_proof #> relabel_isar_proof_nicely #> rationalize_obtains_in_isar_proofs ctxt) in (case (num_chained, add_isar_steps (steps_of_isar_proof isar_proof) 0) of (0, 1) => one_line_proof_text ctxt 0 (if is_less (play_outcome_ord (play_outcome, one_line_play)) then (case isar_proof of Proof {steps = [Prove {facts = (_, gfs), proof_methods = meth :: _, ...}], ...} => let val used_facts' = map_filter (fn s => if exists (fn (s', (sc, _)) => s' = s andalso sc = Chained) used_facts then NONE else SOME (s, (Global, General))) gfs in ((used_facts', (meth, play_outcome)), banner, subgoal, subgoal_count) end) else one_line_params) ^ (if isar_proofs = SOME true then "\n(No Isar proof available)" else "") | (_, num_steps) => let val msg = (if verbose then [string_of_int num_steps ^ " step" ^ plural_s num_steps] else []) @ (if do_preplay then [string_of_play_outcome play_outcome] else []) in one_line_proof_text ctxt 0 one_line_params ^ "\n\nIsar proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^ Active.sendback_markup_command (string_of_isar_proof ctxt subgoal subgoal_count isar_proof) end) end end in if debug then generate_proof_text () else (case try generate_proof_text () of SOME s => s | NONE => one_line_proof_text ctxt 0 one_line_params ^ (if isar_proofs = SOME true then "\nWarning: Isar proof construction failed" else "")) end fun isar_proof_would_be_a_good_idea (_, play) = (case play of Played _ => false | Play_Timed_Out time => time > Time.zeroTime | Play_Failed => true) fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained (one_line_params as ((_, preplay), _, _, _)) = (if isar_proofs = SOME true orelse (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then isar_proof_text ctxt debug num_chained isar_proofs smt_proofs isar_params else one_line_proof_text ctxt num_chained) one_line_params end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML @@ -1,251 +1,255 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML Author: Steffen Juilf Smolka, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Preplaying of Isar proofs. *) signature SLEDGEHAMMER_ISAR_PREPLAY = sig type play_outcome = Sledgehammer_Proof_Methods.play_outcome type proof_method = Sledgehammer_Proof_Methods.proof_method type isar_step = Sledgehammer_Isar_Proof.isar_step type isar_proof = Sledgehammer_Isar_Proof.isar_proof type label = Sledgehammer_Isar_Proof.label val trace : bool Config.T type isar_preplay_data val peek_at_outcome : play_outcome Lazy.lazy -> play_outcome val enrich_context_with_local_facts : isar_proof -> Proof.context -> Proof.context val preplay_isar_step_for_method : Proof.context -> thm list -> Time.time -> proof_method -> isar_step -> play_outcome val preplay_isar_step : Proof.context -> thm list -> Time.time -> proof_method list -> isar_step -> (proof_method * play_outcome) list val set_preplay_outcomes_of_isar_step : Proof.context -> Time.time -> isar_preplay_data Unsynchronized.ref -> isar_step -> (proof_method * play_outcome) list -> unit val forced_intermediate_preplay_outcome_of_isar_step : isar_preplay_data -> label -> play_outcome val preplay_outcome_of_isar_step_for_method : isar_preplay_data -> label -> proof_method -> play_outcome Lazy.lazy val fastest_method_of_isar_step : isar_preplay_data -> label -> proof_method val preplay_outcome_of_isar_proof : isar_preplay_data -> isar_proof -> play_outcome end; structure Sledgehammer_Isar_Preplay : SLEDGEHAMMER_ISAR_PREPLAY = struct open ATP_Proof_Reconstruct open Sledgehammer_Util open Sledgehammer_Proof_Methods open Sledgehammer_Isar_Proof val trace = Attrib.setup_config_bool \<^binding>\sledgehammer_preplay_trace\ (K false) fun peek_at_outcome outcome = if Lazy.is_finished outcome then Lazy.force outcome else Play_Timed_Out Time.zeroTime fun par_list_map_filter_with_timeout _ _ _ _ [] = [] | par_list_map_filter_with_timeout get_time min_timeout timeout0 f xs = let val next_timeout = Unsynchronized.ref timeout0 fun apply_f x = let val timeout = !next_timeout in if timeout <= min_timeout then NONE else let val y = f timeout x in (case get_time y of SOME time => next_timeout := Time.min (time, !next_timeout) | _ => ()); SOME y end end in chop_groups (Multithreading.max_threads ()) xs |> map (map_filter I o Par_List.map apply_f) |> flat end fun enrich_context_with_local_facts proof ctxt = let val thy = Proof_Context.theory_of ctxt fun enrich_with_fact l t = Proof_Context.put_thms false (string_of_label l, SOME [Skip_Proof.make_thm thy t]) val enrich_with_assms = fold (uncurry enrich_with_fact) fun enrich_with_proof (Proof {assumptions, steps = isar_steps, ...}) = enrich_with_assms assumptions #> fold enrich_with_step isar_steps and enrich_with_step (Prove {label, goal, subproofs, ...}) = enrich_with_fact label goal #> fold enrich_with_proof subproofs | enrich_with_step _ = I in enrich_with_proof proof ctxt end fun preplay_trace ctxt assmsp concl outcome = let val ctxt = ctxt |> Config.put show_markup true val assms = op @ assmsp val time = Pretty.str ("[" ^ string_of_play_outcome outcome ^ "]") val assms = Pretty.enum " and " "using " " shows " (map (Thm.pretty_thm ctxt) assms) val concl = Syntax.pretty_term ctxt concl in tracing (Pretty.string_of (Pretty.blk (2, Pretty.breaks [time, assms, concl]))) end fun take_time timeout tac arg = let val timing = Timing.start () in (Timeout.apply timeout tac arg; Played (#cpu (Timing.result timing))) handle Timeout.TIMEOUT _ => Play_Timed_Out timeout end fun resolve_fact_names ctxt names = (names |>> map string_of_label |> apply2 (maps (thms_of_name ctxt))) handle ERROR msg => error ("preplay error: " ^ msg) fun thm_of_proof ctxt (Proof {fixes, assumptions, steps}) = let val thy = Proof_Context.theory_of ctxt val concl = (case try List.last steps of SOME (Prove {obtains = [], goal, ...}) => goal | _ => raise Fail "preplay error: malformed subproof") val var_idx = maxidx_of_term concl + 1 fun var_of_free (x, T) = Var ((x, var_idx), T) val subst = map (`var_of_free #> swap #> apfst Free) fixes in Logic.list_implies (assumptions |> map snd, concl) |> subst_free subst |> Skip_Proof.make_thm thy end (* may throw exceptions *) fun raw_preplay_step_for_method ctxt chained timeout meth (Prove {obtains = xs, goal, subproofs, facts, ...}) = let val goal = (case xs of [] => goal | _ => (* proof obligation: !!thesis. (!!x1...xn. t ==> thesis) ==> thesis (cf. "~~/src/Pure/Isar/obtain.ML") *) let val frees = map Free xs val thesis = Free (singleton (Variable.variant_frees ctxt frees) ("thesis", HOLogic.boolT)) val thesis_prop = HOLogic.mk_Trueprop thesis (* !!x1...xn. t ==> thesis *) val inner_prop = fold_rev Logic.all frees (Logic.mk_implies (goal, thesis_prop)) in (* !!thesis. (!!x1...xn. t ==> thesis) ==> thesis *) Logic.all thesis (Logic.mk_implies (inner_prop, thesis_prop)) end) val assmsp = resolve_fact_names ctxt facts |>> append (map (thm_of_proof ctxt) subproofs) |>> append chained fun prove () = Goal.prove ctxt [] [] goal (fn {context = ctxt, ...} => HEADGOAL (tac_of_proof_method ctxt assmsp meth)) handle ERROR msg => error ("Preplay error: " ^ msg) val play_outcome = take_time timeout prove () in if Config.get ctxt trace then preplay_trace ctxt assmsp goal play_outcome else (); play_outcome end -fun preplay_isar_step_for_method ctxt chained timeout meth = - try (raw_preplay_step_for_method ctxt chained timeout meth) #> the_default Play_Failed +fun preplay_isar_step_for_method ctxt chained timeout meth step = + if timeout = Time.zeroTime then + Play_Timed_Out Time.zeroTime + else + try (raw_preplay_step_for_method ctxt chained timeout meth) step + |> the_default Play_Failed val min_preplay_timeout = seconds 0.002 fun preplay_isar_step ctxt chained timeout0 hopeless step = let fun preplay timeout meth = (meth, preplay_isar_step_for_method ctxt chained timeout meth step) fun get_time (_, Played time) = SOME time | get_time _ = NONE val meths = proof_methods_of_isar_step step |> subtract (op =) hopeless in par_list_map_filter_with_timeout get_time min_preplay_timeout timeout0 preplay meths |> sort (play_outcome_ord o apply2 snd) end type isar_preplay_data = (proof_method * play_outcome Lazy.lazy) list Canonical_Label_Tab.table fun time_of_play (Played time) = time | time_of_play (Play_Timed_Out time) = time fun add_preplay_outcomes Play_Failed _ = Play_Failed | add_preplay_outcomes _ Play_Failed = Play_Failed | add_preplay_outcomes (Played time1) (Played time2) = Played (time1 + time2) | add_preplay_outcomes play1 play2 = Play_Timed_Out (Time.+ (apply2 time_of_play (play1, play2))) fun set_preplay_outcomes_of_isar_step ctxt timeout preplay_data (step as Prove {label = l, proof_methods, ...}) meths_outcomes0 = let fun lazy_preplay meth = Lazy.lazy (fn () => preplay_isar_step_for_method ctxt [] timeout meth step) val meths_outcomes = meths_outcomes0 |> map (apsnd Lazy.value) |> fold (fn meth => AList.default (op =) (meth, lazy_preplay meth)) proof_methods in preplay_data := Canonical_Label_Tab.update (l, fold (AList.update (op =)) meths_outcomes []) (!preplay_data) end | set_preplay_outcomes_of_isar_step _ _ _ _ _ = () fun get_best_method_outcome force = tap (List.app (K () o Lazy.future Future.default_params o snd)) (* could be parallelized *) #> map (apsnd force) #> sort (play_outcome_ord o apply2 snd) #> hd fun forced_intermediate_preplay_outcome_of_isar_step preplay_data l = let val meths_outcomes as (_, outcome1) :: _ = the (Canonical_Label_Tab.lookup preplay_data l) in if forall (not o Lazy.is_finished o snd) meths_outcomes then (case Lazy.force outcome1 of outcome as Played _ => outcome | _ => snd (get_best_method_outcome Lazy.force meths_outcomes)) else let val outcome = snd (get_best_method_outcome peek_at_outcome meths_outcomes) in if outcome = Play_Timed_Out Time.zeroTime then snd (get_best_method_outcome Lazy.force meths_outcomes) else outcome end end fun preplay_outcome_of_isar_step_for_method preplay_data l = AList.lookup (op =) (the (Canonical_Label_Tab.lookup preplay_data l)) #> the_default (Lazy.value (Play_Timed_Out Time.zeroTime)) fun fastest_method_of_isar_step preplay_data = the o Canonical_Label_Tab.lookup preplay_data #> get_best_method_outcome Lazy.force #> fst fun forced_outcome_of_step preplay_data (Prove {label, proof_methods, ...}) = Lazy.force (preplay_outcome_of_isar_step_for_method preplay_data label (hd proof_methods)) | forced_outcome_of_step _ _ = Played Time.zeroTime fun preplay_outcome_of_isar_proof preplay_data (Proof {steps, ...}) = fold_isar_steps (add_preplay_outcomes o forced_outcome_of_step preplay_data) steps (Played Time.zeroTime) end;