diff --git a/src/HOL/Tools/ATP/atp_problem.ML b/src/HOL/Tools/ATP/atp_problem.ML --- a/src/HOL/Tools/ATP/atp_problem.ML +++ b/src/HOL/Tools/ATP/atp_problem.ML @@ -1,995 +1,996 @@ (* Title: HOL/Tools/ATP/atp_problem.ML Author: Jia Meng, Cambridge University Computer Laboratory and NICTA Author: Jasmin Blanchette, TU Muenchen Abstract representation of ATP problems and TPTP syntax. *) signature ATP_PROBLEM = sig datatype ('a, 'b) atp_term = ATerm of ('a * 'b list) * ('a, 'b) atp_term list | AAbs of (('a * 'b) * ('a, 'b) atp_term) * ('a, 'b) atp_term list datatype atp_quantifier = AForall | AExists datatype atp_connective = ANot | AAnd | AOr | AImplies | AIff datatype ('a, 'b, 'c, 'd) atp_formula = ATyQuant of atp_quantifier * ('b * 'd list) list * ('a, 'b, 'c, 'd) atp_formula | AQuant of atp_quantifier * ('a * 'b option) list * ('a, 'b, 'c, 'd) atp_formula | AConn of atp_connective * ('a, 'b, 'c, 'd) atp_formula list | AAtom of 'c datatype 'a atp_type = AType of ('a * 'a list) * 'a atp_type list | AFun of 'a atp_type * 'a atp_type | APi of 'a list * 'a atp_type type term_order = {is_lpo : bool, gen_weights : bool, gen_prec : bool, gen_simp : bool} datatype polymorphism = Monomorphic | Polymorphic datatype thf_choice = THF_Predicate_Free | THF_Without_Choice | THF_With_Choice datatype atp_format = CNF | CNF_UEQ | FOF | TFF of polymorphism | THF of polymorphism * thf_choice | DFG of polymorphism datatype atp_formula_role = Axiom | Definition | Lemma | Hypothesis | Conjecture | Negated_Conjecture | Plain | Type_Role | Unknown datatype 'a atp_problem_line = Class_Decl of string * 'a * 'a list | Type_Decl of string * 'a * int | Sym_Decl of string * 'a * 'a atp_type | Datatype_Decl of string * ('a * 'a list) list * 'a atp_type * ('a, 'a atp_type) atp_term list * bool | Class_Memb of string * ('a * 'a list) list * 'a atp_type * 'a | Formula of (string * string) * atp_formula_role * ('a, 'a atp_type, ('a, 'a atp_type) atp_term, 'a) atp_formula * (string, string atp_type) atp_term option * (string, string atp_type) atp_term list type 'a atp_problem = (string * 'a atp_problem_line list) list val tptp_cnf : string val tptp_fof : string val tptp_tff : string val tptp_thf : string val tptp_has_type : string val tptp_type_of_types : string val tptp_bool_type : string val tptp_individual_type : string val tptp_fun_type : string val tptp_product_type : string val tptp_forall : string val tptp_ho_forall : string val tptp_pi_binder : string val tptp_exists : string val tptp_ho_exists : string val tptp_choice : string val tptp_ho_choice : string val tptp_hilbert_choice : string val tptp_hilbert_the : string val tptp_not : string val tptp_and : string val tptp_not_and : string val tptp_or : string val tptp_not_or : string val tptp_implies : string val tptp_if : string val tptp_iff : string val tptp_not_iff : string val tptp_app : string val tptp_not_infix : string val tptp_equal : string val tptp_not_equal : string val tptp_old_equal : string val tptp_false : string val tptp_true : string val tptp_lambda : string val tptp_empty_list : string val dfg_individual_type : string val isabelle_info_prefix : string val isabelle_info : bool -> string -> int -> (string, 'a) atp_term list val extract_isabelle_status : (string, 'a) atp_term list -> string option val extract_isabelle_rank : (string, 'a) atp_term list -> int val inductionN : string val introN : string val inductiveN : string val elimN : string val simpN : string val non_rec_defN : string val rec_defN : string val rankN : string val minimum_rank : int val default_rank : int val default_term_order_weight : int val is_tptp_equal : string -> bool val is_built_in_tptp_symbol : string -> bool val is_tptp_variable : string -> bool val is_tptp_user_symbol : string -> bool val bool_atype : (string * string) atp_type val individual_atype : (string * string) atp_type val mk_anot : ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'c, 'd) atp_formula val mk_aconn : atp_connective -> ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'c, 'd) atp_formula val mk_app : (string, 'a) atp_term -> (string, 'a) atp_term -> (string, 'a) atp_term val mk_apps : (string, 'a) atp_term -> (string, 'a) atp_term list -> (string, 'a) atp_term val mk_simple_aterm: 'a -> ('a, 'b) atp_term val aconn_fold : bool option -> (bool option -> 'a -> 'b -> 'b) -> atp_connective * 'a list -> 'b -> 'b val aconn_map : bool option -> (bool option -> 'a -> ('b, 'c, 'd, 'e) atp_formula) -> atp_connective * 'a list -> ('b, 'c, 'd, 'e) atp_formula val formula_fold : bool option -> (bool option -> 'c -> 'e -> 'e) -> ('a, 'b, 'c, 'd) atp_formula -> 'e -> 'e val formula_map : ('c -> 'e) -> ('a, 'b, 'c, 'd) atp_formula -> ('a, 'b, 'e, 'd) atp_formula val strip_atype : 'a atp_type -> 'a list * ('a atp_type list * 'a atp_type) val is_format_higher_order : atp_format -> bool val tptp_string_of_format : atp_format -> string + val tptp_string_of_role : atp_formula_role -> string val tptp_string_of_line : atp_format -> string atp_problem_line -> string val lines_of_atp_problem : atp_format -> term_order -> (unit -> (string * int) list) -> string atp_problem -> string list val ensure_cnf_problem : (string * string) atp_problem -> (string * string) atp_problem val filter_cnf_ueq_problem : (string * string) atp_problem -> (string * string) atp_problem val declared_in_atp_problem : 'a atp_problem -> ('a list * 'a list) * 'a list val nice_atp_problem : bool -> atp_format -> ('a * (string * string) atp_problem_line list) list -> ('a * string atp_problem_line list) list * (string Symtab.table * string Symtab.table) option end; structure ATP_Problem : ATP_PROBLEM = struct open ATP_Util val parens = enclose "(" ")" (** ATP problem **) datatype ('a, 'b) atp_term = ATerm of ('a * 'b list) * ('a, 'b) atp_term list | AAbs of (('a * 'b) * ('a, 'b) atp_term) * ('a, 'b) atp_term list datatype atp_quantifier = AForall | AExists datatype atp_connective = ANot | AAnd | AOr | AImplies | AIff datatype ('a, 'b, 'c, 'd) atp_formula = ATyQuant of atp_quantifier * ('b * 'd list) list * ('a, 'b, 'c, 'd) atp_formula | AQuant of atp_quantifier * ('a * 'b option) list * ('a, 'b, 'c, 'd) atp_formula | AConn of atp_connective * ('a, 'b, 'c, 'd) atp_formula list | AAtom of 'c datatype 'a atp_type = AType of ('a * 'a list) * 'a atp_type list | AFun of 'a atp_type * 'a atp_type | APi of 'a list * 'a atp_type type term_order = {is_lpo : bool, gen_weights : bool, gen_prec : bool, gen_simp : bool} datatype polymorphism = Monomorphic | Polymorphic datatype thf_choice = THF_Predicate_Free | THF_Without_Choice | THF_With_Choice datatype atp_format = CNF | CNF_UEQ | FOF | TFF of polymorphism | THF of polymorphism * thf_choice | DFG of polymorphism datatype atp_formula_role = Axiom | Definition | Lemma | Hypothesis | Conjecture | Negated_Conjecture | Plain | Type_Role | Unknown datatype 'a atp_problem_line = Class_Decl of string * 'a * 'a list | Type_Decl of string * 'a * int | Sym_Decl of string * 'a * 'a atp_type | Datatype_Decl of string * ('a * 'a list) list * 'a atp_type * ('a, 'a atp_type) atp_term list * bool | Class_Memb of string * ('a * 'a list) list * 'a atp_type * 'a | Formula of (string * string) * atp_formula_role * ('a, 'a atp_type, ('a, 'a atp_type) atp_term, 'a) atp_formula * (string, string atp_type) atp_term option * (string, string atp_type) atp_term list type 'a atp_problem = (string * 'a atp_problem_line list) list (* official TPTP syntax *) val tptp_cnf = "cnf" val tptp_fof = "fof" val tptp_tff = "tff" val tptp_thf = "thf" val tptp_has_type = ":" val tptp_type_of_types = "$tType" val tptp_bool_type = "$o" val tptp_individual_type = "$i" val tptp_fun_type = ">" val tptp_product_type = "*" val tptp_forall = "!" val tptp_ho_forall = "!!" val tptp_pi_binder = "!>" val tptp_exists = "?" val tptp_ho_exists = "??" val tptp_choice = "@+" val tptp_ho_choice = "@@+" val tptp_not = "~" val tptp_and = "&" val tptp_not_and = "~&" val tptp_or = "|" val tptp_not_or = "~|" val tptp_implies = "=>" val tptp_if = "<=" val tptp_iff = "<=>" val tptp_not_iff = "<~>" val tptp_app = "@" val tptp_hilbert_choice = "@+" val tptp_hilbert_the = "@-" val tptp_not_infix = "!" val tptp_equal = "=" val tptp_not_equal = "!=" val tptp_old_equal = "equal" val tptp_false = "$false" val tptp_true = "$true" val tptp_lambda = "^" val tptp_empty_list = "[]" val dfg_individual_type = "iii" (* cannot clash *) val isabelle_info_prefix = "isabelle_" val inductionN = "induction" val introN = "intro" val inductiveN = "inductive" val elimN = "elim" val simpN = "simp" val non_rec_defN = "non_rec_def" val rec_defN = "rec_def" val rankN = "rank" val minimum_rank = 0 val default_rank = 1000 val default_term_order_weight = 1 (* Currently, only SPASS 3.8ds and (to a lesser extent) Metis can process Isabelle metainformation. *) fun isabelle_info generate_info status rank = if generate_info then [] |> rank <> default_rank ? cons (ATerm ((isabelle_info_prefix ^ rankN, []), [ATerm ((string_of_int rank, []), [])])) |> status <> "" ? cons (ATerm ((isabelle_info_prefix ^ status, []), [])) else [] fun extract_isabelle_status (ATerm ((s, []), []) :: _) = try (unprefix isabelle_info_prefix) s | extract_isabelle_status _ = NONE fun extract_isabelle_rank (tms as _ :: _) = (case List.last tms of ATerm ((_, []), [ATerm ((rank, []), [])]) => the (Int.fromString rank) | _ => default_rank) | extract_isabelle_rank _ = default_rank fun is_tptp_equal s = (s = tptp_equal orelse s = tptp_old_equal) fun is_built_in_tptp_symbol s = s = tptp_old_equal orelse not (Char.isAlpha (String.sub (s, 0))) fun is_tptp_variable s = s <> "" andalso Char.isUpper (String.sub (s, 0)) val is_tptp_user_symbol = not o (is_tptp_variable orf is_built_in_tptp_symbol) val bool_atype = AType ((`I tptp_bool_type, []), []) val individual_atype = AType ((`I tptp_individual_type, []), []) fun raw_polarities_of_conn ANot = (SOME false, NONE) | raw_polarities_of_conn AAnd = (SOME true, SOME true) | raw_polarities_of_conn AOr = (SOME true, SOME true) | raw_polarities_of_conn AImplies = (SOME false, SOME true) | raw_polarities_of_conn AIff = (NONE, NONE) fun polarities_of_conn NONE = K (NONE, NONE) | polarities_of_conn (SOME pos) = raw_polarities_of_conn #> not pos ? apply2 (Option.map not) fun mk_anot (AConn (ANot, [phi])) = phi | mk_anot phi = AConn (ANot, [phi]) fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2]) fun mk_app t u = ATerm ((tptp_app, []), [t, u]) fun mk_apps f xs = fold (fn x => fn f => mk_app f x) xs f fun mk_simple_aterm p = ATerm ((p, []), []) fun aconn_fold pos f (ANot, [phi]) = f (Option.map not pos) phi | aconn_fold pos f (AImplies, [phi1, phi2]) = f (Option.map not pos) phi1 #> f pos phi2 | aconn_fold pos f (AAnd, phis) = fold (f pos) phis | aconn_fold pos f (AOr, phis) = fold (f pos) phis | aconn_fold _ f (_, phis) = fold (f NONE) phis fun aconn_map pos f (ANot, [phi]) = AConn (ANot, [f (Option.map not pos) phi]) | aconn_map pos f (AImplies, [phi1, phi2]) = AConn (AImplies, [f (Option.map not pos) phi1, f pos phi2]) | aconn_map pos f (AAnd, phis) = AConn (AAnd, map (f pos) phis) | aconn_map pos f (AOr, phis) = AConn (AOr, map (f pos) phis) | aconn_map _ f (c, phis) = AConn (c, map (f NONE) phis) fun formula_fold pos f = let fun fld pos (AQuant (_, _, phi)) = fld pos phi | fld pos (ATyQuant (_, _, phi)) = fld pos phi | fld pos (AConn conn) = aconn_fold pos fld conn | fld pos (AAtom tm) = f pos tm in fld pos end fun formula_map f (AQuant (q, xs, phi)) = AQuant (q, xs, formula_map f phi) | formula_map f (ATyQuant (q, xs, phi)) = ATyQuant (q, xs, formula_map f phi) | formula_map f (AConn (c, phis)) = AConn (c, map (formula_map f) phis) | formula_map f (AAtom tm) = AAtom (f tm) fun strip_api (APi (tys, ty)) = strip_api ty |>> append tys | strip_api ty = ([], ty) fun strip_afun (AFun (ty1, ty2)) = strip_afun ty2 |>> cons ty1 | strip_afun ty = ([], ty) fun strip_atype ty = ty |> strip_api ||> strip_afun fun is_function_atype ty = snd (snd (strip_atype ty)) <> AType ((tptp_bool_type, []), []) fun is_predicate_atype ty = not (is_function_atype ty) fun is_nontrivial_predicate_atype (AType _) = false | is_nontrivial_predicate_atype ty = is_predicate_atype ty fun is_format_higher_order (THF _) = true | is_format_higher_order _ = false fun is_format_typed (TFF _) = true | is_format_typed (THF _) = true | is_format_typed (DFG _) = true | is_format_typed _ = false fun tptp_string_of_role Axiom = "axiom" | tptp_string_of_role Definition = "definition" | tptp_string_of_role Lemma = "lemma" | tptp_string_of_role Hypothesis = "hypothesis" | tptp_string_of_role Conjecture = "conjecture" | tptp_string_of_role Negated_Conjecture = "negated_conjecture" | tptp_string_of_role Plain = "plain" | tptp_string_of_role Type_Role = "type" | tptp_string_of_role Unknown = "unknown" fun tptp_string_of_app _ func [] = func | tptp_string_of_app format func args = if is_format_higher_order format then "(" ^ space_implode (" " ^ tptp_app ^ " ") (func :: args) ^ ")" else func ^ "(" ^ commas args ^ ")" fun uncurry_type (APi (tys, ty)) = APi (tys, uncurry_type ty) | uncurry_type (ty as AFun (ty1 as AType _, ty2)) = (case uncurry_type ty2 of AFun (ty' as AType ((s, _), tys), ty) => AFun (AType ((tptp_product_type, []), ty1 :: (if s = tptp_product_type then tys else [ty'])), ty) | _ => ty) | uncurry_type (ty as AType _) = ty | uncurry_type _ = raise Fail "unexpected higher-order type in first-order format" val suffix_type_of_types = suffix (" " ^ tptp_has_type ^ " " ^ tptp_type_of_types) fun str_of_type format ty = let val dfg = (case format of DFG _ => true | _ => false) fun str _ (AType ((s, _), [])) = if dfg andalso s = tptp_individual_type then dfg_individual_type else s | str rhs (AType ((s, _), tys)) = if s = tptp_fun_type then let val [ty1, ty2] = tys in str rhs (AFun (ty1, ty2)) end else let val ss = tys |> map (str false) in if s = tptp_product_type then ss |> space_implode (if dfg then ", " else " " ^ tptp_product_type ^ " ") |> (not dfg andalso length ss > 1) ? parens else tptp_string_of_app format s ss end | str rhs (AFun (ty1, ty2)) = (str false ty1 |> dfg ? parens) ^ " " ^ (if dfg then "" else tptp_fun_type ^ " ") ^ str true ty2 |> not rhs ? parens | str _ (APi (ss, ty)) = if dfg then "[" ^ commas ss ^ "], " ^ str true ty else tptp_pi_binder ^ "[" ^ commas (map suffix_type_of_types ss) ^ "]: " ^ str false ty in str true ty end fun string_of_type (format as THF _) ty = str_of_type format ty | string_of_type format ty = str_of_type format (uncurry_type ty) fun tptp_string_of_quantifier AForall = tptp_forall | tptp_string_of_quantifier AExists = tptp_exists fun tptp_string_of_connective ANot = tptp_not | tptp_string_of_connective AAnd = tptp_and | tptp_string_of_connective AOr = tptp_or | tptp_string_of_connective AImplies = tptp_implies | tptp_string_of_connective AIff = tptp_iff fun string_of_bound_var format (s, ty) = s ^ (if is_format_typed format then " " ^ tptp_has_type ^ " " ^ (ty |> the_default (AType ((tptp_individual_type, []), [])) |> string_of_type format) else "") fun tptp_string_of_term _ (ATerm ((s, []), [])) = s | tptp_string_of_term format (ATerm ((s, tys), ts)) = let val of_type = string_of_type format val of_term = tptp_string_of_term format fun app () = tptp_string_of_app format s (map (of_type #> is_format_higher_order format ? parens) tys @ map of_term ts) in if s = tptp_empty_list then (* used for lists in the optional "source" field of a derivation *) "[" ^ commas (map of_term ts) ^ "]" else if is_tptp_equal s then space_implode (" " ^ tptp_equal ^ " ") (map of_term ts) |> is_format_higher_order format ? parens else if s = tptp_ho_forall orelse s = tptp_ho_exists then (case ts of [AAbs (((s', ty), tm), [])] => (* Use syntactic sugar "!" and "?" instead of "!!" and "??" whenever possible, to work around LEO-II 1.2.8 parser limitation. *) tptp_string_of_formula format (AQuant (if s = tptp_ho_forall then AForall else AExists, [(s', SOME ty)], AAtom tm)) | _ => app ()) else if s = tptp_choice then (case ts of [AAbs (((s', ty), tm), args)] => (* There is code in "ATP_Problem_Generate" to ensure that "Eps" is always applied to an abstraction. *) tptp_string_of_app format (tptp_choice ^ "[" ^ s' ^ " : " ^ of_type ty ^ "]: " ^ of_term tm ^ "" |> parens) (map of_term args) | _ => app ()) else if s = tptp_not then (* agsyHOL doesn't like negations applied like this: "~ @ t". *) (case ts of [t] => s ^ " " ^ (of_term t |> parens) |> parens | _ => s) else if member (op =) [tptp_and, tptp_or, tptp_implies, tptp_if, tptp_iff, tptp_not_iff, tptp_equal] s then (* agsyHOL doesn't like connectives applied like this: "& @ t1 @ t2". *) (case ts of [t1, t2] => (of_term t1 |> parens) ^ " " ^ s ^ " " ^ (of_term t2 |> parens) |> parens | _ => app ()) else app () end | tptp_string_of_term (format as THF _) (AAbs (((s, ty), tm), args)) = tptp_string_of_app format ("(^[" ^ s ^ " : " ^ string_of_type format ty ^ "]: " ^ tptp_string_of_term format tm ^ ")") (map (tptp_string_of_term format) args) | tptp_string_of_term _ _ = raise Fail "unexpected term in first-order format" and tptp_string_of_formula format (ATyQuant (q, xs, phi)) = tptp_string_of_quantifier q ^ "[" ^ commas (map (suffix_type_of_types o string_of_type format o fst) xs) ^ "]: " ^ tptp_string_of_formula format phi |> parens | tptp_string_of_formula format (AQuant (q, xs, phi)) = tptp_string_of_quantifier q ^ "[" ^ commas (map (string_of_bound_var format) xs) ^ "]: " ^ tptp_string_of_formula format phi |> parens | tptp_string_of_formula format (AConn (ANot, [AAtom (ATerm (("=" (* tptp_equal *), []), ts))])) = space_implode (" " ^ tptp_not_infix ^ tptp_equal ^ " ") (map (tptp_string_of_term format) ts) |> is_format_higher_order format ? parens | tptp_string_of_formula format (AConn (c, [phi])) = tptp_string_of_connective c ^ " " ^ (tptp_string_of_formula format phi |> is_format_higher_order format ? parens) |> parens | tptp_string_of_formula format (AConn (c, phis)) = space_implode (" " ^ tptp_string_of_connective c ^ " ") (map (tptp_string_of_formula format) phis) |> parens | tptp_string_of_formula format (AAtom tm) = tptp_string_of_term format tm fun tptp_string_of_format CNF = tptp_cnf | tptp_string_of_format CNF_UEQ = tptp_cnf | tptp_string_of_format FOF = tptp_fof | tptp_string_of_format (TFF _) = tptp_tff | tptp_string_of_format (THF _) = tptp_thf | tptp_string_of_format (DFG _) = raise Fail "non-TPTP format" val atype_of_types = AType ((tptp_type_of_types, []), []) fun nary_type_decl_type n = funpow n (curry AFun atype_of_types) atype_of_types fun maybe_alt "" = "" | maybe_alt s = " % " ^ s fun tptp_string_of_line format (Type_Decl (ident, ty, ary)) = tptp_string_of_line format (Sym_Decl (ident, ty, nary_type_decl_type ary)) | tptp_string_of_line format (Sym_Decl (ident, sym, ty)) = tptp_string_of_format format ^ "(" ^ ident ^ ", type,\n " ^ sym ^ " : " ^ string_of_type format ty ^ ").\n" | tptp_string_of_line format (Formula ((ident, alt), kind, phi, source, info)) = tptp_string_of_format format ^ "(" ^ ident ^ ", " ^ tptp_string_of_role kind ^ "," ^ "\n (" ^ tptp_string_of_formula format phi ^ ")" ^ (case source of SOME tm => ", " ^ tptp_string_of_term format tm | NONE => if null info then "" else ", []") ^ (case info of [] => "" | tms => ", [" ^ commas (map (tptp_string_of_term format) tms) ^ "]") ^ ")." ^ maybe_alt alt ^ "\n" fun tptp_lines format = maps (fn (_, []) => [] | (heading, lines) => "\n% " ^ heading ^ " (" ^ string_of_int (length lines) ^ ")\n" :: map (tptp_string_of_line format) lines) fun arity_of_type (APi (tys, ty)) = arity_of_type ty |>> Integer.add (length tys) | arity_of_type (AFun (_, ty)) = arity_of_type ty ||> Integer.add 1 | arity_of_type _ = (0, 0) fun string_of_arity (0, n) = string_of_int n | string_of_arity (m, n) = string_of_int m ^ "+" ^ string_of_int n val dfg_class_inter = space_implode " & " fun dfg_string_of_term (ATerm ((s, tys), tms)) = s ^ (if null tys then "" else "<" ^ commas (map (string_of_type (DFG Polymorphic)) tys) ^ ">") ^ (if null tms then "" else "(" ^ commas (map dfg_string_of_term tms) ^ ")") | dfg_string_of_term _ = raise Fail "unexpected atom in first-order format" fun dfg_string_of_formula poly gen_simp info = let val str_of_typ = string_of_type (DFG poly) fun str_of_bound_typ (ty, []) = str_of_typ ty | str_of_bound_typ (ty, cls) = str_of_typ ty ^ " : " ^ dfg_class_inter cls fun suffix_tag top_level s = if top_level then (case extract_isabelle_status info of SOME s' => if s' = non_rec_defN then s ^ ":lt" else if (s' = simpN orelse s' = rec_defN) andalso gen_simp then s ^ ":lr" else s | NONE => s) else s fun str_of_atom top_level (ATerm ((s, tys), tms)) = let val s' = if is_tptp_equal s then "equal" |> suffix_tag top_level else if s = tptp_true then "true" else if s = tptp_false then "false" else s in dfg_string_of_term (ATerm ((s', tys), tms)) end | str_of_atom _ _ = raise Fail "unexpected atom in first-order format" fun str_of_quant AForall = "forall" | str_of_quant AExists = "exists" fun str_of_conn _ ANot = "not" | str_of_conn _ AAnd = "and" | str_of_conn _ AOr = "or" | str_of_conn _ AImplies = "implies" | str_of_conn top_level AIff = "equiv" |> suffix_tag top_level fun str_of_formula top_level (ATyQuant (q, xs, phi)) = str_of_quant q ^ "_sorts([" ^ commas (map str_of_bound_typ xs) ^ "], " ^ str_of_formula top_level phi ^ ")" | str_of_formula top_level (AQuant (q, xs, phi)) = str_of_quant q ^ "([" ^ commas (map (string_of_bound_var (DFG poly)) xs) ^ "], " ^ str_of_formula top_level phi ^ ")" | str_of_formula top_level (AConn (c, phis)) = str_of_conn top_level c ^ "(" ^ commas (map (str_of_formula false) phis) ^ ")" | str_of_formula top_level (AAtom tm) = str_of_atom top_level tm in str_of_formula true end fun maybe_enclose bef aft "" = "% " ^ bef ^ aft | maybe_enclose bef aft s = bef ^ s ^ aft fun dfg_lines poly {is_lpo, gen_weights, gen_prec, gen_simp} ord_info problem = let val typ = string_of_type (DFG poly) val term = dfg_string_of_term fun spair (s, s') = "(" ^ s ^ ", " ^ s' ^ ")" fun tm_ary sym ty = spair (sym, string_of_arity (arity_of_type ty)) fun ty_ary 0 ty = ty | ty_ary n ty = "(" ^ ty ^ ", " ^ string_of_int n ^ ")" fun fun_typ sym ty = "function(" ^ sym ^ ", " ^ typ ty ^ ")." fun pred_typ sym ty = let val (ty_vars, (tys, _)) = strip_atype ty |>> (fn [] => [] | xs => ["[" ^ commas xs ^ "]"]) in "predicate(" ^ commas (sym :: ty_vars @ map typ tys) ^ ")." end fun bound_tvar (ty, []) = ty | bound_tvar (ty, cls) = ty ^ " : " ^ dfg_class_inter cls fun binder_typ xs ty = (if null xs then "" else "[" ^ commas (map bound_tvar xs) ^ "], ") ^ typ ty fun sort_decl xs ty cl = "sort(" ^ binder_typ xs ty ^ ", " ^ cl ^ ")." fun datatype_decl xs ty tms exhaust = "datatype(" ^ commas (binder_typ xs ty :: map term tms @ (if exhaust then [] else ["..."])) ^ ")." fun subclass_of sub super = "subclass(" ^ sub ^ ", " ^ super ^ ")." fun formula pred (Formula ((ident, alt), kind, phi, _, info)) = if pred kind then let val rank = extract_isabelle_rank info in "formula(" ^ dfg_string_of_formula poly gen_simp info phi ^ ", " ^ ident ^ (if rank = default_rank then "" else ", " ^ string_of_int rank) ^ ")." ^ maybe_alt alt |> SOME end else NONE | formula _ _ = NONE fun filt f = problem |> map (map_filter f o snd) |> filter_out null val func_aries = filt (fn Sym_Decl (_, sym, ty) => if is_function_atype ty then SOME (tm_ary sym ty) else NONE | _ => NONE) |> flat |> commas |> maybe_enclose "functions [" "]." val pred_aries = filt (fn Sym_Decl (_, sym, ty) => if is_predicate_atype ty then SOME (tm_ary sym ty) else NONE | _ => NONE) |> flat |> commas |> maybe_enclose "predicates [" "]." val sorts = filt (try (fn Type_Decl (_, ty, ary) => ty_ary ary ty)) @ [[ty_ary 0 dfg_individual_type]] |> flat |> commas |> maybe_enclose "sorts [" "]." val classes = filt (try (fn Class_Decl (_, cl, _) => cl)) |> flat |> commas |> maybe_enclose "classes [" "]." val ord_info = if gen_weights orelse gen_prec then ord_info () else [] val do_term_order_weights = (if gen_weights then ord_info else []) |> map (spair o apsnd string_of_int) |> commas |> maybe_enclose "weights [" "]." val syms = [func_aries, pred_aries, do_term_order_weights, sorts, classes] val func_decls = filt (fn Sym_Decl (_, sym, ty) => if is_function_atype ty then SOME (fun_typ sym ty) else NONE | _ => NONE) |> flat val pred_decls = filt (fn Sym_Decl (_, sym, ty) => if is_nontrivial_predicate_atype ty then SOME (pred_typ sym ty) else NONE | _ => NONE) |> flat val datatype_decls = filt (try (fn Datatype_Decl (_, xs, ty, tms, exhaust) => datatype_decl xs ty tms exhaust)) |> flat val sort_decls = filt (try (fn Class_Memb (_, xs, ty, cl) => sort_decl xs ty cl)) |> flat val subclass_decls = filt (try (fn Class_Decl (_, sub, supers) => map (subclass_of sub) supers)) |> flat |> flat val decls = func_decls @ pred_decls @ datatype_decls @ sort_decls @ subclass_decls val axioms = filt (formula (curry (op <>) Conjecture)) |> separate [""] |> flat val conjs = filt (formula (curry (op =) Conjecture)) |> separate [""] |> flat val settings = (if is_lpo then ["set_flag(Ordering, 1)."] else []) @ (if gen_prec then [ord_info |> map fst |> rev |> commas |> maybe_enclose "set_precedence(" ")."] else []) fun list_of _ [] = [] | list_of heading ss = "list_of_" ^ heading ^ ".\n" :: map (suffix "\n") ss @ ["end_of_list.\n\n"] in "\nbegin_problem(isabelle).\n\n" :: list_of "descriptions" ["name({**}).", "author({**}).", "status(unknown).", "description({**})."] @ list_of "symbols" syms @ list_of "declarations" decls @ list_of "formulae(axioms)" axioms @ list_of "formulae(conjectures)" conjs @ list_of "settings(SPASS)" settings @ ["end_problem.\n"] end fun lines_of_atp_problem format ord ord_info problem = "% This file was generated by Isabelle (most likely Sledgehammer)\n\ \% " ^ timestamp () ^ "\n" :: (case format of DFG poly => dfg_lines poly ord ord_info | _ => tptp_lines format) problem (** CNF (Metis) and CNF UEQ (Waldmeister) **) fun is_line_negated (Formula (_, _, AConn (ANot, _), _, _)) = true | is_line_negated _ = false fun is_line_cnf_ueq (Formula (_, _, AAtom (ATerm (((s, _), _), _)), _, _)) = is_tptp_equal s | is_line_cnf_ueq _ = false fun open_conjecture_term (ATerm (((s, s'), tys), tms)) = ATerm ((if is_tptp_variable s then (s |> Name.desymbolize (SOME false), s') else (s, s'), tys), tms |> map open_conjecture_term) | open_conjecture_term _ = raise Fail "unexpected higher-order term" fun open_formula conj = let (* We are conveniently assuming that all bound variable names are distinct, which should be the case for the formulas we generate. *) fun opn (pos as SOME true) (AQuant (AForall, _, phi)) = opn pos phi | opn (pos as SOME false) (AQuant (AExists, _, phi)) = opn pos phi | opn pos (AConn (ANot, [phi])) = mk_anot (opn (Option.map not pos) phi) | opn pos (AConn (c, [phi1, phi2])) = let val (pos1, pos2) = polarities_of_conn pos c in AConn (c, [opn pos1 phi1, opn pos2 phi2]) end | opn _ (AAtom t) = AAtom (t |> conj ? open_conjecture_term) | opn _ phi = phi in opn (SOME (not conj)) end fun open_formula_line (Formula (ident, kind, phi, source, info)) = Formula (ident, kind, open_formula (kind = Conjecture) phi, source, info) | open_formula_line line = line fun negate_conjecture_line (Formula (ident, Conjecture, phi, source, info)) = Formula (ident, Hypothesis, mk_anot phi, source, info) | negate_conjecture_line line = line exception CLAUSIFY of unit (* This "clausification" only expands syntactic sugar, such as "phi => psi" to "~ phi | psi" and "phi <=> psi" to "~ phi | psi" and "~ psi | phi". We don't attempt to distribute conjunctions over disjunctions. *) fun clausify_formula pos (phi as AAtom _) = [phi |> not pos ? mk_anot] | clausify_formula pos (AConn (ANot, [phi])) = clausify_formula (not pos) phi | clausify_formula true (AConn (AOr, [phi1, phi2])) = (phi1, phi2) |> apply2 (clausify_formula true) |> uncurry (map_product (mk_aconn AOr)) | clausify_formula false (AConn (AAnd, [phi1, phi2])) = (phi1, phi2) |> apply2 (clausify_formula false) |> uncurry (map_product (mk_aconn AOr)) | clausify_formula true (AConn (AImplies, [phi1, phi2])) = clausify_formula true (AConn (AOr, [mk_anot phi1, phi2])) | clausify_formula true (AConn (AIff, phis)) = clausify_formula true (AConn (AImplies, phis)) @ clausify_formula true (AConn (AImplies, rev phis)) | clausify_formula _ _ = raise CLAUSIFY () fun clausify_formula_line (Formula ((ident, alt), kind, phi, source, info)) = let val (n, phis) = phi |> try (clausify_formula true) |> these |> `length in map2 (fn phi => fn j => Formula ((ident ^ replicate_string (j - 1) "x", alt), kind, phi, source, info)) phis (1 upto n) end | clausify_formula_line _ = [] fun ensure_cnf_line line = line |> open_formula_line |> negate_conjecture_line |> clausify_formula_line fun ensure_cnf_problem problem = problem |> map (apsnd (maps ensure_cnf_line)) fun filter_cnf_ueq_problem problem = problem |> map (apsnd (map open_formula_line #> filter is_line_cnf_ueq #> map negate_conjecture_line)) |> (fn problem => let val lines = problem |> maps snd val conjs = lines |> filter is_line_negated in if length conjs = 1 andalso conjs <> lines then problem else [] end) (** Symbol declarations **) fun add_declared_in_line (Class_Decl (_, cl, _)) = apfst (apfst (cons cl)) | add_declared_in_line (Type_Decl (_, ty, _)) = apfst (apsnd (cons ty)) | add_declared_in_line (Sym_Decl (_, sym, _)) = apsnd (cons sym) | add_declared_in_line _ = I fun declared_in_atp_problem problem = fold (fold add_declared_in_line o snd) problem (([], []), []) (** Nice names **) val no_qualifiers = let fun skip [] = [] | skip (#"." :: cs) = skip cs | skip (c :: cs) = if Char.isAlphaNum c then skip cs else c :: keep cs and keep [] = [] | keep (#"." :: cs) = skip cs | keep (c :: cs) = c :: keep cs in String.explode #> rev #> keep #> rev #> String.implode end (* Long names can slow down the ATPs. *) val max_readable_name_size = 20 (* "equal" is reserved by some ATPs. "op" is also reserved, to avoid the unreadable "op_1", "op_2", etc., in the problem files. "eq" is reserved to ensure that "HOL.eq" is correctly mapped to equality (not clear whether this is still necessary). *) val reserved_nice_names = [tptp_old_equal, "op", "eq"] (* hack to get the same hashing across Mirabelle runs (see "mirabelle.pl") *) fun cleanup_mirabelle_name s = let val mirabelle_infix = "_Mirabelle_" val random_suffix_len = 10 val (s1, s2) = Substring.position mirabelle_infix (Substring.full s) in if Substring.isEmpty s2 then s else Substring.string s1 ^ Substring.string (Substring.triml (size mirabelle_infix + random_suffix_len) s2) end fun readable_name protect full_name s = (if s = full_name then s else s |> no_qualifiers |> unquote_tvar |> Name.desymbolize (SOME (Char.isUpper (String.sub (full_name, 0)))) |> (fn s => if size s > max_readable_name_size then String.substring (s, 0, max_readable_name_size div 2 - 4) ^ string_of_int (hash_string (cleanup_mirabelle_name full_name)) ^ String.extract (s, size s - max_readable_name_size div 2 + 4, NONE) else s) |> (fn s => if member (op =) reserved_nice_names s then full_name else s)) |> protect fun nice_name _ (full_name, _) NONE = (full_name, NONE) | nice_name protect (full_name, desired_name) (SOME the_pool) = if is_built_in_tptp_symbol full_name then (full_name, SOME the_pool) else (case Symtab.lookup (fst the_pool) full_name of SOME nice_name => (nice_name, SOME the_pool) | NONE => let val nice_prefix = readable_name protect full_name desired_name fun add j = let val nice_name = nice_prefix ^ (if j = 1 then "" else string_of_int j) in (case Symtab.lookup (snd the_pool) nice_name of SOME full_name' => if full_name = full_name' then (nice_name, the_pool) else add (j + 1) | NONE => (nice_name, (Symtab.update_new (full_name, nice_name) (fst the_pool), Symtab.update_new (nice_name, full_name) (snd the_pool)))) end in add 1 |> apsnd SOME end) fun avoid_clash_with_alt_ergo_type_vars s = if is_tptp_variable s then s else s ^ "_" fun avoid_clash_with_dfg_keywords s = let val n = String.size s in if n < 2 orelse (n = 2 andalso String.sub (s, 0) = String.sub (s, 1)) orelse String.isSubstring "_" s then s else if is_tptp_variable s then s ^ "_" else String.substring (s, 0, n - 1) ^ String.str (Char.toUpper (String.sub (s, n - 1))) end fun nice_atp_problem readable_names format problem = let val empty_pool = if readable_names then SOME (Symtab.empty, Symtab.empty) else NONE val avoid_clash = (case format of TFF Polymorphic => avoid_clash_with_alt_ergo_type_vars | DFG _ => avoid_clash_with_dfg_keywords | _ => I) val nice_name = nice_name avoid_clash fun nice_bound_tvars xs = fold_map (nice_name o fst) xs ##>> fold_map (fold_map nice_name o snd) xs #>> op ~~ fun nice_type (AType ((name, clss), tys)) = nice_name name ##>> fold_map nice_name clss ##>> fold_map nice_type tys #>> AType | nice_type (AFun (ty1, ty2)) = nice_type ty1 ##>> nice_type ty2 #>> AFun | nice_type (APi (names, ty)) = fold_map nice_name names ##>> nice_type ty #>> APi fun nice_term (ATerm ((name, tys), ts)) = nice_name name ##>> fold_map nice_type tys ##>> fold_map nice_term ts #>> ATerm | nice_term (AAbs (((name, ty), tm), args)) = nice_name name ##>> nice_type ty ##>> nice_term tm ##>> fold_map nice_term args #>> AAbs fun nice_formula (ATyQuant (q, xs, phi)) = fold_map (nice_type o fst) xs ##>> fold_map (fold_map nice_name o snd) xs ##>> nice_formula phi #>> (fn ((tys, cls), phi) => ATyQuant (q, tys ~~ cls, phi)) | nice_formula (AQuant (q, xs, phi)) = fold_map (nice_name o fst) xs ##>> fold_map (fn (_, NONE) => pair NONE | (_, SOME ty) => nice_type ty #>> SOME) xs ##>> nice_formula phi #>> (fn ((ss, ts), phi) => AQuant (q, ss ~~ ts, phi)) | nice_formula (AConn (c, phis)) = fold_map nice_formula phis #>> curry AConn c | nice_formula (AAtom tm) = nice_term tm #>> AAtom fun nice_line (Class_Decl (ident, cl, cls)) = nice_name cl ##>> fold_map nice_name cls #>> (fn (cl, cls) => Class_Decl (ident, cl, cls)) | nice_line (Type_Decl (ident, ty, ary)) = nice_name ty #>> (fn ty => Type_Decl (ident, ty, ary)) | nice_line (Sym_Decl (ident, sym, ty)) = nice_name sym ##>> nice_type ty #>> (fn (sym, ty) => Sym_Decl (ident, sym, ty)) | nice_line (Datatype_Decl (ident, xs, ty, tms, exhaust)) = nice_bound_tvars xs ##>> nice_type ty ##>> fold_map nice_term tms #>> (fn ((xs, ty), tms) => Datatype_Decl (ident, xs, ty, tms, exhaust)) | nice_line (Class_Memb (ident, xs, ty, cl)) = nice_bound_tvars xs ##>> nice_type ty ##>> nice_name cl #>> (fn ((xs, ty), cl) => Class_Memb (ident, xs, ty, cl)) | nice_line (Formula (ident, kind, phi, source, info)) = nice_formula phi #>> (fn phi => Formula (ident, kind, phi, source, info)) fun nice_problem problem = fold_map (fn (heading, lines) => fold_map nice_line lines #>> pair heading) problem in nice_problem problem empty_pool end end; 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,725 +1,741 @@ (* 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 | CantConnect | TimedOut | Inappropriate | OutOfResources | NoPerl | NoLibwwwPerl | 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 (* Named ATPs *) val agsyholN : string val alt_ergoN : string val eN : string val e_parN : string val iproverN : string val leo2N : string val leo3N : string val pirateN : string val satallaxN : string val snarkN : string val spassN : string val vampireN : string val waldmeisterN : string val z3_tptpN : string val zipperpositionN : string val remote_prefix : string val agsyhol_core_rule : string val spass_input_rule : string val spass_pre_skolemize_rule : string val spass_skolemize_rule : string val z3_tptp_core_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_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_thf_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 -> 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 end; structure ATP_Proof : ATP_PROOF = struct open ATP_Util open ATP_Problem (* Named ATPs *) val agsyholN = "agsyhol" val alt_ergoN = "alt_ergo" val eN = "e" val e_parN = "e_par" val iproverN = "iprover" val leo2N = "leo2" val leo3N = "leo3" val pirateN = "pirate" val satallaxN = "satallax" val snarkN = "snark" val spassN = "spass" val vampireN = "vampire" val waldmeisterN = "waldmeister" val z3_tptpN = "z3_tptp" val zipperpositionN = "zipperposition" val remote_prefix = "remote_" 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 z3_tptp_core_rule = "__z3_tptp_core" (* arbitrary *) exception UNRECOGNIZED_ATP_PROOF of unit datatype atp_failure = MaybeUnprovable | Unprovable | GaveUp | ProofMissing | ProofIncomplete | ProofUnparsable | UnsoundProof of bool * string list | CantConnect | TimedOut | Inappropriate | OutOfResources | NoPerl | NoLibwwwPerl | 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 "" val missing_message_tail = " appears to be missing; you will need to install it if you want to invoke \ \remote provers" fun from_lemmas [] = "" | from_lemmas ss = " from " ^ space_implode " " (Try.serial_commas "and" (map quote ss)) fun string_of_atp_failure MaybeUnprovable = "The generated problem is maybe unprovable" | string_of_atp_failure Unprovable = "The generated problem is unprovable" | string_of_atp_failure GaveUp = "The prover gave up" | string_of_atp_failure ProofMissing = "The prover claims the conjecture is a theorem but did not provide a proof" | string_of_atp_failure ProofIncomplete = "The prover claims the conjecture is a theorem but provided an incomplete proof" | string_of_atp_failure ProofUnparsable = "The prover claims the conjecture is a theorem but provided an unparsable proof" | string_of_atp_failure (UnsoundProof (false, ss)) = "The prover derived \"False\"" ^ from_lemmas ss ^ "; specify a sound type encoding or omit the \"type_enc\" option" | string_of_atp_failure (UnsoundProof (true, ss)) = "The prover derived \"False\"" ^ from_lemmas ss ^ ", which could be due to a bug in Sledgehammer or to inconsistent axioms (including \"sorry\"s)" | string_of_atp_failure CantConnect = "Cannot connect to server" | string_of_atp_failure TimedOut = "Timed out" | string_of_atp_failure Inappropriate = "The generated problem lies outside the prover's scope" | string_of_atp_failure OutOfResources = "The prover ran out of resources" | string_of_atp_failure NoPerl = "Perl" ^ missing_message_tail | string_of_atp_failure NoLibwwwPerl = "The Perl module \"libwww-perl\"" ^ missing_message_tail | string_of_atp_failure MalformedInput = "The generated problem is malformed" | string_of_atp_failure MalformedOutput = "The prover output is malformed" | string_of_atp_failure Interrupted = "The prover was interrupted" | string_of_atp_failure Crashed = "The prover crashed" | string_of_atp_failure InternalError = "An internal prover error occurred" | string_of_atp_failure (UnknownError s) = "A prover error occurred" ^ (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 = (case (extract_tstplike_proof proof_delims output, extract_known_atp_failure known_failures output) 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)) 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. *) (case apply2 (Int.fromString o perhaps (try (unprefix vampire_fact_prefix))) q of (SOME i, SOME j) => int_ord (i, j) | _ => raise Fail "not Vampire") 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 (for "z3_tptp"). *) 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_inference = Inference_Source ("", []) 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 || parse_inference_source >> Inference_Source || parse_introduced_source >> Introduced_Source || scan_general_id >> (fn s => Inference_Source ("", [s])) (* for E *) || skip_term >> K dummy_inference) 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_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_arg x = ($$ "(" |-- parse_term --| $$ ")" --| parse_optional_type_signature || scan_general_id -- parse_optional_type_signature -- Scan.optional ($$ "<" |-- parse_types --| $$ ">") [] -- Scan.optional ($$ "(" |-- parse_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_term x = (parse_arg -- Scan.repeat ($$ tptp_app |-- parse_arg) --| parse_optional_type_signature >> list_app) x and parse_terms x = (parse_term ::: Scan.repeat ($$ "," |-- parse_term)) x fun parse_atom x = (parse_term -- Scan.option (Scan.option ($$ tptp_not_infix) --| $$ tptp_equal -- parse_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_literal x = ((Scan.repeat ($$ tptp_not) >> length) -- ($$ "(" |-- parse_formula --| $$ ")" || parse_quantified_formula || parse_atom) >> (fn (n, phi) => phi |> n mod 2 = 1 ? mk_anot)) x and parse_formula x = (parse_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_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_quantified_formula x = (($$ tptp_forall >> K AForall || $$ tptp_exists >> K AExists) --| $$ "[" -- parse_terms --| $$ "]" --| $$ ":" -- parse_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)) dummy_inference 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 val tptp_binary_ops = [tptp_and, tptp_not_and, tptp_or, tptp_not_or, tptp_implies, tptp_if, tptp_iff, tptp_not_iff, tptp_equal, tptp_not_equal, tptp_app] fun parse_one_in_list xs = foldl1 (op ||) (map Scan.this_string xs) fun parse_binary_op x = (parse_one_in_list tptp_binary_ops >> (fn c => if c = tptp_equal then "equal" else c)) x val parse_fo_quantifier = parse_one_in_list [tptp_forall, tptp_exists, tptp_lambda, tptp_hilbert_choice, tptp_hilbert_the] val parse_ho_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_thf_app (ATerm ((x, ty), arg)) = if x = tptp_app then (case arg of ATerm ((x, ty), arg) :: t => remove_thf_app (ATerm ((x, ty), map remove_thf_app arg @ t)) | [AAbs ((var, tvar), phi), t] => remove_thf_app (AAbs ((var, tvar), map remove_thf_app phi @ [t]))) else ATerm ((x, ty), map remove_thf_app arg) | remove_thf_app (AAbs (((x, ty), arg), t)) = AAbs (((x, ty), remove_thf_app arg), t) fun parse_typed_var x = (Scan.repeat (scan_general_id -- Scan.option ($$ tptp_has_type |-- parse_type) --| Scan.option (Scan.this_string ",")) || $$ "(" |-- parse_typed_var --| $$ ")") x fun parse_simple_thf_term x = (parse_fo_quantifier -- ($$ "[" |-- parse_typed_var --| $$ "]" --| $$ ":") -- parse_thf_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_thf_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_ho_quantifier >> mk_simple_aterm || $$ "(" |-- parse_thf_term --| $$ ")" || parse_binary_op >> mk_simple_aterm) x and parse_thf_term x = (parse_simple_thf_term -- Scan.option (parse_binary_op -- parse_thf_term) >> (fn (t1, SOME (c, t2)) => if c = tptp_app then mk_app t1 t2 else mk_apps (mk_simple_aterm c) [t1, t2] | (t, NONE) => t)) x fun parse_thf_formula x = (parse_thf_term #>> remove_thf_app #>> AAtom) x fun parse_tstp_thf_line problem = (Scan.this_string tptp_thf -- $$ "(") |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ "," -- (parse_thf_formula || skip_term >> K dummy_phi) -- parse_tstp_extra_arguments --| $$ ")" --| $$ "." >> (fn (((num, role), phi), deps) => let val role' = role_of_tptp_string role val ((name, phi), rule, deps) = (case deps of File_Source (_, SOME s) => if role' = Definition then (((num, map fst (find_formula_in_problem phi problem)), phi), "", []) else (((num, [s]), phi), "", []) | Inference_Source (rule, deps) => (((num, []), phi), rule, deps)) in [(name, role', phi, rule, map (rpair []) deps)] end) (* Syntax: (cnf|fof|tff|thf)\(, , \). The could be an identifier, but we assume integers. *) fun parse_tstp_line problem = ((Scan.this_string tptp_cnf || Scan.this_string tptp_fof || Scan.this_string tptp_tff || Scan.this_string tptp_thf) -- $$ "(") |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ "," -- (parse_formula || 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 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, "", []) | File_Source _ => (((num, map fst (find_formula_in_problem phi problem)), phi), role, "", []) | Inference_Source (rule, deps) => (((num, []), phi), role, rule, deps) | Introduced_Source rule => (((num, []), phi), Lemma, rule, [])) 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) (**** 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_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 parse_pirate_dependency x = (Scan.option ($$ "-") |-- scan_general_id) x fun parse_pirate_dependencies x = Scan.repeat (parse_pirate_dependency --| Scan.option ($$ "," || $$ " ")) x fun parse_pirate_file_source x = ((Scan.this_string "Input" || Scan.this_string "Conj") |-- $$ "(" |-- scan_general_id --| $$ ")") x fun parse_pirate_inference_source x = (scan_general_id -- ($$ "(" |-- parse_pirate_dependencies --| $$ ")")) x fun parse_pirate_source x = (parse_pirate_file_source >> (fn s => File_Source ("", SOME s)) || parse_pirate_inference_source >> Inference_Source) x (* Syntax: || -> . origin\(\) *) fun parse_pirate_line x = (scan_general_id --| Scan.repeat (~$$ "|") -- parse_horn_clause --| $$ "." --| Scan.this_string "origin" --| $$ "(" -- parse_pirate_source --| $$ ")" >> (fn ((((num, u), source))) => let val (names, rule, deps) = (case source of File_Source (_, SOME s) => ([s], spass_input_rule, []) | Inference_Source (rule, deps) => ([], rule, deps)) in [((num, names), Unknown, u, rule, map (rpair []) deps)] end)) x fun core_inference inf fact = ((fact, [fact]), Unknown, dummy_phi, inf, []) (* Syntax: SZS core ... *) fun parse_z3_tptp_core_line x = (Scan.this_string "SZS core" |-- Scan.repeat ($$ " " |-- scan_general_id) >> map (core_inference z3_tptp_core_rule)) x fun parse_line local_name problem = (* Satallax is handled separately, in "atp_satallax.ML". *) if local_name = leo2N orelse local_name = leo3N then parse_tstp_thf_line problem else if local_name = spassN then parse_spass_line else if local_name = pirateN then parse_pirate_line else if local_name = z3_tptpN then parse_z3_tptp_core_line else parse_tstp_line 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 + end; diff --git a/src/HOL/Tools/SMT/smt_solver.ML b/src/HOL/Tools/SMT/smt_solver.ML --- a/src/HOL/Tools/SMT/smt_solver.ML +++ b/src/HOL/Tools/SMT/smt_solver.ML @@ -1,315 +1,315 @@ (* Title: HOL/Tools/SMT/smt_solver.ML Author: Sascha Boehme, TU Muenchen SMT solvers registry and SMT tactic. *) signature SMT_SOLVER = sig (*configuration*) datatype outcome = Unsat | Sat | Unknown | Time_Out type parsed_proof = {outcome: SMT_Failure.failure option, fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option, atp_proof: unit -> (term, string) ATP_Proof.atp_step list} type solver_config = {name: string, class: Proof.context -> SMT_Util.class, avail: unit -> bool, command: unit -> string list, options: Proof.context -> string list, smt_options: (string * string) list, default_max_relevant: int, outcome: string -> string list -> outcome * string list, parse_proof: (Proof.context -> SMT_Translate.replay_data -> ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list -> parsed_proof) option, replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option} (*registry*) val add_solver: solver_config -> theory -> theory val default_max_relevant: Proof.context -> string -> int (*filter*) val smt_filter: Proof.context -> thm -> ((string * ATP_Problem_Generate.stature) * thm) list -> int -> Time.time -> parsed_proof (*tactic*) val smt_tac: Proof.context -> thm list -> int -> tactic val smt_tac': Proof.context -> thm list -> int -> tactic end; structure SMT_Solver: SMT_SOLVER = struct (* interface to external solvers *) local fun make_command command options problem_path proof_path = Bash.strings (command () @ options) ^ " " ^ File.bash_platform_path problem_path ^ " > " ^ File.bash_path proof_path ^ " 2>&1" fun with_trace ctxt msg f x = let val _ = SMT_Config.trace_msg ctxt (fn () => msg) () in f x end fun run ctxt name mk_cmd input = (case SMT_Config.certificates_of ctxt of NONE => if not (SMT_Config.is_available ctxt name) then error ("The SMT solver " ^ quote name ^ " is not installed") else if Config.get ctxt SMT_Config.debug_files = "" then with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input else let val base_path = Path.explode (Config.get ctxt SMT_Config.debug_files) val in_path = Path.ext "smt_in" base_path val out_path = Path.ext "smt_out" base_path in Cache_IO.raw_run mk_cmd input in_path out_path end | SOME certs => (case Cache_IO.lookup certs input of (NONE, key) => if Config.get ctxt SMT_Config.read_only_certificates then error ("Bad certificate cache: missing certificate") else Cache_IO.run_and_cache certs key mk_cmd input | (SOME output, _) => with_trace ctxt ("Using cached certificate from " ^ Path.print (Cache_IO.cache_path_of certs) ^ " ...") I output)) (* Z3 returns 1 if "get-proof" or "get-model" fails. veriT returns 255. *) val normal_return_codes = [0, 1, 255] fun run_solver ctxt name mk_cmd input = let fun pretty tag lines = Pretty.string_of (Pretty.big_list tag (map Pretty.str lines)) val _ = SMT_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input val ({elapsed, ...}, {redirected_output = res, output = err, return_code}) = Timing.timing (SMT_Config.with_timeout ctxt (run ctxt name mk_cmd)) input val _ = SMT_Config.trace_msg ctxt (pretty "Solver:") err val output = drop_suffix (equal "") res val _ = SMT_Config.trace_msg ctxt (pretty "Result:") output val _ = SMT_Config.trace_msg ctxt (pretty "Time (ms):") [\<^make_string> (Time.toMilliseconds elapsed)] val _ = SMT_Config.statistics_msg ctxt (pretty "Time (ms):") [\<^make_string> (Time.toMilliseconds elapsed)] val _ = member (op =) normal_return_codes return_code orelse raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code) in output end fun trace_assms ctxt = SMT_Config.trace_msg ctxt (Pretty.string_of o Pretty.big_list "Assertions:" o map (Thm.pretty_thm ctxt o snd)) fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT_Translate.replay_data) = let fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p] fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T) fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t) in SMT_Config.trace_msg ctxt (fn () => Pretty.string_of (Pretty.big_list "Names:" [ Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)), Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) () end in fun invoke name command smt_options ithms ctxt = let val options = SMT_Config.solver_options_of ctxt val comments = [space_implode " " options] val (str, replay_data as {context = ctxt', ...}) = ithms |> tap (trace_assms ctxt) |> SMT_Translate.translate ctxt smt_options comments ||> tap trace_replay_data in (run_solver ctxt' name (make_command command options) str, replay_data) end end (* configuration *) datatype outcome = Unsat | Sat | Unknown | Time_Out type parsed_proof = {outcome: SMT_Failure.failure option, fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option, atp_proof: unit -> (term, string) ATP_Proof.atp_step list} type solver_config = {name: string, class: Proof.context -> SMT_Util.class, avail: unit -> bool, command: unit -> string list, options: Proof.context -> string list, smt_options: (string * string) list, default_max_relevant: int, outcome: string -> string list -> outcome * string list, parse_proof: (Proof.context -> SMT_Translate.replay_data -> ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list -> parsed_proof) option, replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option} (* check well-sortedness *) val has_topsort = Term.exists_type (Term.exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)) (* top sorts cause problems with atomization *) fun check_topsort ctxt thm = if has_topsort (Thm.prop_of thm) then (SMT_Normalize.drop_fact_warning ctxt thm; TrueI) else thm (* registry *) type solver_info = { command: unit -> string list, smt_options: (string * string) list, default_max_relevant: int, parse_proof: Proof.context -> SMT_Translate.replay_data -> ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list -> parsed_proof, replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm} structure Solvers = Generic_Data ( type T = solver_info Symtab.table val empty = Symtab.empty val extend = I fun merge data = Symtab.merge (K true) data ) local fun parse_proof outcome parse_proof0 outer_ctxt replay_data xfacts prems concl output = (case outcome output of (Unsat, lines) => (case parse_proof0 of SOME pp => pp outer_ctxt replay_data xfacts prems concl lines | NONE => {outcome = NONE, fact_ids = NONE, atp_proof = K []}) | (Time_Out, _) => raise SMT_Failure.SMT (SMT_Failure.Time_Out) | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat))) fun replay outcome replay0 oracle outer_ctxt (replay_data as {context = ctxt, ...} : SMT_Translate.replay_data) output = (case outcome output of (Unsat, lines) => if Config.get ctxt SMT_Config.oracle then oracle () else (case replay0 of SOME replay => replay outer_ctxt replay_data lines | NONE => error "No proof reconstruction for solver -- \ \declare [[smt_oracle]] to allow oracle") | (Time_Out, _) => raise SMT_Failure.SMT (SMT_Failure.Time_Out) | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat))) val cfalse = Thm.cterm_of \<^context> \<^prop>\False\ in fun add_solver ({name, class, avail, command, options, smt_options, default_max_relevant, outcome, parse_proof = parse_proof0, replay = replay0} : solver_config) = let fun solver oracle = { command = command, smt_options = smt_options, default_max_relevant = default_max_relevant, parse_proof = parse_proof (outcome name) parse_proof0, replay = replay (outcome name) replay0 oracle} val info = {name = name, class = class, avail = avail, options = options} in Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, oracle) => Context.theory_map (Solvers.map (Symtab.update_new (name, solver oracle)))) #> Context.theory_map (SMT_Config.add_solver info) end end fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name) fun name_and_info_of ctxt = let val name = SMT_Config.solver_of ctxt in (name, get_info ctxt name) end val default_max_relevant = #default_max_relevant oo get_info fun apply_solver_and_replay ctxt thms0 = let val thms = map (check_topsort ctxt) thms0 val (name, {command, smt_options, replay, ...}) = name_and_info_of ctxt val (output, replay_data) = invoke name command smt_options (SMT_Normalize.normalize ctxt thms) ctxt in replay ctxt replay_data output end -(* filter *) +(* filter (for Sledgehammer) *) fun smt_filter ctxt0 goal xfacts i time_limit = let val ctxt = ctxt0 |> Config.put SMT_Config.timeout (Time.toReal time_limit) val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i NONE goal fun negate ct = Thm.dest_comb ct ||> Thm.apply \<^cterm>\Not\ |-> Thm.apply val cprop = (case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt concl)) of SOME ct => ct | NONE => raise SMT_Failure.SMT (SMT_Failure.Other_Failure "cannot atomize goal")) val conjecture = Thm.assume cprop val facts = map snd xfacts val thms = conjecture :: prems @ facts val thms' = map (check_topsort ctxt) thms val (name, {command, smt_options, parse_proof, ...}) = name_and_info_of ctxt val (output, replay_data) = invoke name command smt_options (SMT_Normalize.normalize ctxt thms') ctxt in parse_proof ctxt replay_data xfacts (map Thm.prop_of prems) (Thm.term_of concl) output end handle SMT_Failure.SMT fail => {outcome = SOME fail, fact_ids = NONE, atp_proof = K []} (* SMT tactic *) local fun str_of ctxt fail = "Solver " ^ SMT_Config.solver_of ctxt ^ ": " ^ SMT_Failure.string_of_failure fail fun safe_solve ctxt facts = SOME (apply_solver_and_replay ctxt facts) handle SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) => (SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE) | SMT_Failure.SMT (fail as SMT_Failure.Time_Out) => error ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^ SMT_Failure.string_of_failure fail ^ " (setting the " ^ "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)") | SMT_Failure.SMT fail => error (str_of ctxt fail) fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1 | resolve _ NONE = no_tac fun tac prove ctxt rules = CONVERSION (SMT_Normalize.atomize_conv ctxt) THEN' resolve_tac ctxt @{thms ccontr} THEN' SUBPROOF (fn {context = ctxt', prems, ...} => resolve ctxt' (prove ctxt' (rules @ prems))) ctxt in val smt_tac = tac safe_solve val smt_tac' = tac (SOME oo apply_solver_and_replay) end end; diff --git a/src/HOL/Tools/SMT/z3_isar.ML b/src/HOL/Tools/SMT/z3_isar.ML --- a/src/HOL/Tools/SMT/z3_isar.ML +++ b/src/HOL/Tools/SMT/z3_isar.ML @@ -1,120 +1,120 @@ (* Title: HOL/Tools/SMT/z3_isar.ML Author: Jasmin Blanchette, TU Muenchen Z3 proofs as generic ATP proofs for Isar proof reconstruction. *) signature Z3_ISAR = sig val atp_proof_of_z3_proof: Proof.context -> term list -> thm list -> term list -> term -> (string * term) list -> int list -> int -> (int * string) list -> Z3_Proof.z3_step list -> (term, string) ATP_Proof.atp_step list end; structure Z3_Isar: Z3_ISAR = struct open ATP_Util open ATP_Problem open ATP_Proof open ATP_Proof_Reconstruct open SMTLIB_Isar val z3_apply_def_rule = Z3_Proof.string_of_rule Z3_Proof.Apply_Def val z3_hypothesis_rule = Z3_Proof.string_of_rule Z3_Proof.Hypothesis val z3_intro_def_rule = Z3_Proof.string_of_rule Z3_Proof.Intro_Def val z3_lemma_rule = Z3_Proof.string_of_rule Z3_Proof.Lemma fun inline_z3_defs _ [] = [] | inline_z3_defs defs ((name, role, t, rule, deps) :: lines) = if rule = z3_intro_def_rule then let val def = t |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> swap in inline_z3_defs (insert (op =) def defs) (map (replace_dependencies_in_line (name, [])) lines) end else if rule = z3_apply_def_rule then inline_z3_defs defs (map (replace_dependencies_in_line (name, [])) lines) else (name, role, Term.subst_atomic defs t, rule, deps) :: inline_z3_defs defs lines fun add_z3_hypotheses [] = I | add_z3_hypotheses hyps = HOLogic.dest_Trueprop #> curry s_imp (Library.foldr1 s_conj (map HOLogic.dest_Trueprop hyps)) #> HOLogic.mk_Trueprop fun inline_z3_hypotheses _ _ [] = [] | inline_z3_hypotheses hyp_names hyps ((name, role, t, rule, deps) :: lines) = if rule = z3_hypothesis_rule then inline_z3_hypotheses (name :: hyp_names) (AList.map_default (op =) (t, []) (cons name) hyps) lines else let val deps' = subtract (op =) hyp_names deps in if rule = z3_lemma_rule then (name, role, t, rule, deps') :: inline_z3_hypotheses hyp_names hyps lines else let val add_hyps = filter_out (null o inter (op =) deps o snd) hyps val t' = add_z3_hypotheses (map fst add_hyps) t val hyps' = fold (AList.update (op =) o apsnd (insert (op =) name)) add_hyps hyps in (name, role, t', rule, deps') :: inline_z3_hypotheses hyp_names hyps' lines end end fun dest_alls (Const (\<^const_name>\Pure.all\, _) $ Abs (abs as (_, T, _))) = let val (s', t') = Term.dest_abs abs in dest_alls t' |>> cons (s', T) end | dest_alls t = ([], t) val reorder_foralls = dest_alls #>> sort_by fst #-> fold_rev (Logic.all o Free); fun atp_proof_of_z3_proof ctxt ll_defs rewrite_rules hyp_ts concl_t fact_helper_ts prem_ids conjecture_id fact_helper_ids proof = let - fun steps_of (Z3_Proof.Z3_Step {id, rule, prems, concl, ...}) = + fun steps_of (Z3_Proof.Z3_Step {id, rule, prems, concl, ...}) : (term, string) ATP_Proof.atp_step list = let val sid = string_of_int id val concl' = concl |> reorder_foralls (* crucial for skolemization steps *) |> postprocess_step_conclusion ctxt rewrite_rules ll_defs fun standard_step role = ((sid, []), role, concl', Z3_Proof.string_of_rule rule, map (fn id => (string_of_int id, [])) prems) in (case rule of Z3_Proof.Asserted => let val ss = the_list (AList.lookup (op =) fact_helper_ids id) in (case distinguish_conjecture_and_hypothesis ss id conjecture_id prem_ids fact_helper_ts hyp_ts concl_t of - NONE => [] + NONE => [] (* useless nontheory tautology *) | SOME (role0, concl00) => let val name0 = (sid ^ "a", ss) val concl0 = unskolemize_names ctxt concl00 in (if role0 = Axiom then [] else [(name0, role0, concl0, Z3_Proof.string_of_rule rule, [])]) @ [((sid, []), Plain, concl', Z3_Proof.string_of_rule Z3_Proof.Rewrite, name0 :: normalizing_prems ctxt concl0)] end) end | Z3_Proof.Rewrite => [standard_step Lemma] | Z3_Proof.Rewrite_Star => [standard_step Lemma] | Z3_Proof.Skolemize => [standard_step Lemma] | Z3_Proof.Th_Lemma _ => [standard_step Lemma] | _ => [standard_step Plain]) end in proof |> maps steps_of |> inline_z3_defs [] |> inline_z3_hypotheses [] [] 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,478 +1,485 @@ (* 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 pirate_datatype_rule = "DT" 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_tmp_skolemize_rule = "tmp_skolemize" 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, veriT_tmp_skolemize_rule, z3_skolemize_rule, zipperposition_cnf_rule] 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 val is_datatype_rule = String.isPrefix pirate_datatype_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 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 (is_datatype_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 datatype_methods = [Simp_Method, Simp_Size_Method] 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 :: 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 atp_proof = - fold_rev add_line_pass1 atp_proof0 [] + 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) + |> (fn x => fold_rev add_line_pass1 x []) |> 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 ((l, t), rule) ctxt = let val (skos, meths) = (if is_skolemize_rule rule then (skolems_of ctxt t, skolem_methods) else if is_arith_rule rule then ([], arith_methods) else ([], rewrite_methods)) ||> massage_methods in (Prove ([], skos, l, t, [], ([], []), meths, ""), 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 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 (_, _, _, _, subs, (ls, _), _, _)) = member (op =) ls l orelse exists (is_referenced_in_proof l) subs and is_referenced_in_proof l (Proof (_, _, steps)) = exists (is_referenced_in_step l) steps fun insert_lemma_in_step lem (step as Prove (qs, fix, l, t, subs, (ls, gs), meths, 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') subs in if length (filter I refs) = 1 then let val subs' = map2 (fn false => I | true => insert_lemma_in_proof lem) refs subs in [Prove (qs, fix, l, t, subs', (ls, gs), meths, comment)] end 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 (fix, assms, steps)) = Proof (fix, assms, 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 (if outer then [Show] else [], [], no_label, concl_t, [], sort_facts (the_list predecessor, []), massage_methods systematic_methods', "")) 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 if is_datatype_rule rule then datatype_methods else systematic_methods') |> massage_methods fun prove sub facts = Prove (maybe_show outer c, [], l, t, sub, facts, meths, "") 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 then let val skos = skolems_of ctxt (prop_of_clause g) val subproof = Proof (skos, [], 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 ([], skos, l, t, [], deps, meths, "")) 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 (maybe_show outer c, [], l, t, map isar_case (filter_out (null o snd) cases), sort_facts (the_list predecessor, []), massage_methods systematic_methods', "") in isar_steps outer (SOME l) (step :: accum) infs end and isar_proof outer fix assms lems infs = Proof (fix, assms, fold_rev insert_lemma_in_steps lems (isar_steps outer NONE [] infs)) - val trace = Config.get ctxt trace - 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 ctxt [] 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 (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)) ||> (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 (_, _, [Prove (_, _, _, _, _, (_, gfs), 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 (meth, play) = +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;