diff --git a/src/HOL/Tools/ATP/atp_util.ML b/src/HOL/Tools/ATP/atp_util.ML --- a/src/HOL/Tools/ATP/atp_util.ML +++ b/src/HOL/Tools/ATP/atp_util.ML @@ -1,423 +1,424 @@ (* Title: HOL/Tools/ATP/atp_util.ML Author: Jasmin Blanchette, TU Muenchen General-purpose functions used by the ATP module. *) signature ATP_UTIL = sig val timestamp : unit -> string val hashw : word * word -> word val hashw_string : string * word -> word val hash_string : string -> int val chunk_list : int -> 'a list -> 'a list list val stringN_of_int : int -> int -> string val strip_spaces : bool -> (char -> bool) -> string -> string val strip_spaces_except_between_idents : string -> string val elide_string : int -> string -> string val find_enclosed : string -> string -> string -> string list val nat_subscript : int -> string val unquote_tvar : string -> string val maybe_quote : Keyword.keywords -> string -> string val string_of_ext_time : bool * Time.time -> string val string_of_time : Time.time -> string val type_instance : theory -> typ -> typ -> bool val type_generalization : theory -> typ -> typ -> bool val type_intersect : theory -> typ -> typ -> bool val type_equiv : theory -> typ * typ -> bool val varify_type : Proof.context -> typ -> typ val instantiate_type : theory -> typ -> typ -> typ -> typ val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ val is_type_surely_finite : Proof.context -> typ -> bool val is_type_surely_infinite : Proof.context -> bool -> typ list -> typ -> bool val s_not : term -> term val s_conj : term * term -> term val s_disj : term * term -> term val s_imp : term * term -> term val s_iff : term * term -> term val close_form : term -> term val hol_close_form_prefix : string val hol_close_form : term -> term val hol_open_form : (string -> string) -> term -> term val eta_expand : typ list -> term -> int -> term val cong_extensionalize_term : Proof.context -> term -> term val abs_extensionalize_term : Proof.context -> term -> term val unextensionalize_def : term -> term val transform_elim_prop : term -> term val specialize_type : theory -> (string * typ) -> term -> term val strip_subgoal : thm -> int -> Proof.context -> (string * typ) list * term list * term val extract_lambda_def : (term -> string * typ) -> term -> string * term val short_thm_name : Proof.context -> thm -> string val map_prod : ('a -> 'b) -> ('c -> 'd) -> 'a * 'c -> 'b * 'd val compare_length_with : 'a list -> int -> order end; structure ATP_Util : ATP_UTIL = struct fun timestamp_format time = Date.fmt "%Y-%m-%d %H:%M:%S." (Date.fromTimeLocal time) ^ (StringCvt.padLeft #"0" 3 (string_of_int (Time.toMilliseconds time - 1000 * Time.toSeconds time))) val timestamp = timestamp_format o Time.now (* This hash function is recommended in "Compilers: Principles, Techniques, and Tools" by Aho, Sethi, and Ullman. The "hashpjw" function, which they particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *) fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w)) fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w) fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s fun hash_string s = Word.toInt (hashw_string (s, 0w0)) fun chunk_list _ [] = [] | chunk_list k xs = let val (xs1, xs2) = chop k xs in xs1 :: chunk_list k xs2 end fun stringN_of_int 0 _ = "" | stringN_of_int k n = stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10) fun is_spaceish c = Char.isSpace c orelse c = #"\127" (* DEL -- no idea where these come from *) fun strip_spaces skip_comments is_evil = let fun strip_c_style_comment [] accum = accum | strip_c_style_comment (#"*" :: #"/" :: cs) accum = strip_spaces_in_list true cs accum | strip_c_style_comment (_ :: cs) accum = strip_c_style_comment cs accum and strip_spaces_in_list _ [] accum = accum | strip_spaces_in_list true (#"%" :: cs) accum = strip_spaces_in_list true (cs |> drop_prefix (not_equal #"\n")) accum | strip_spaces_in_list true (#"/" :: #"*" :: cs) accum = strip_c_style_comment cs accum | strip_spaces_in_list _ [c1] accum = accum |> not (is_spaceish c1) ? cons c1 | strip_spaces_in_list skip_comments (cs as [_, _]) accum = accum |> fold (strip_spaces_in_list skip_comments o single) cs | strip_spaces_in_list skip_comments (c1 :: c2 :: c3 :: cs) accum = if is_spaceish c1 then strip_spaces_in_list skip_comments (c2 :: c3 :: cs) accum else if is_spaceish c2 then if is_spaceish c3 then strip_spaces_in_list skip_comments (c1 :: c3 :: cs) accum else strip_spaces_in_list skip_comments (c3 :: cs) (c1 :: accum |> forall is_evil [c1, c3] ? cons #" ") else strip_spaces_in_list skip_comments (c2 :: c3 :: cs) (cons c1 accum) in String.explode #> rpair [] #-> strip_spaces_in_list skip_comments #> rev #> String.implode end fun is_ident_char c = Char.isAlphaNum c orelse c = #"_" val strip_spaces_except_between_idents = strip_spaces true is_ident_char fun elide_string threshold s = if size s > threshold then String.extract (s, 0, SOME (threshold div 2 - 5)) ^ " ...... " ^ String.extract (s, size s - (threshold + 1) div 2 + 6, NONE) else s fun find_enclosed left right s = case first_field left s of SOME (_, s) => (case first_field right s of SOME (enclosed, s) => enclosed :: find_enclosed left right s | NONE => []) | NONE => [] val subscript = implode o map (prefix "\<^sub>") o raw_explode (* FIXME Symbol.explode (?) *) fun nat_subscript n = n |> string_of_int |> not (print_mode_active Print_Mode.ASCII) ? subscript val unquote_tvar = perhaps (try (unprefix "'")) val unquery_var = perhaps (try (unprefix "?")) val is_long_identifier = forall Symbol_Pos.is_identifier o Long_Name.explode fun maybe_quote keywords y = let val s = YXML.content_of y in y |> ((not (is_long_identifier (unquote_tvar s)) andalso not (is_long_identifier (unquery_var s))) orelse Keyword.is_literal keywords s) ? quote end fun string_of_ext_time (plus, time) = let val us = Time.toMicroseconds time in (if plus then "> " else "") ^ (if us < 1000 then string_of_real (Real.fromInt (us div 100) / 10.0) ^ " ms" else if us < 1000000 then signed_string_of_int (us div 1000) ^ " ms" else string_of_real (Real.fromInt (us div 100000) / 10.0) ^ " s") end val string_of_time = string_of_ext_time o pair false fun type_instance thy T T' = Sign.typ_instance thy (T, T') fun type_generalization thy T T' = Sign.typ_instance thy (T', T) fun type_intersect _ (TVar _) _ = true | type_intersect _ _ (TVar _) = true | type_intersect thy T T' = let val tvars = Term.add_tvar_namesT T [] val tvars' = Term.add_tvar_namesT T' [] val maxidx' = maxidx_of_typ T' val T = T |> exists (member (op =) tvars') tvars ? Logic.incr_tvar (maxidx' + 1) val maxidx = Integer.max (maxidx_of_typ T) maxidx' in can (Sign.typ_unify thy (T, T')) (Vartab.empty, maxidx) end val type_equiv = Sign.typ_equiv fun varify_type ctxt T = Variable.polymorphic_types ctxt [Const (\<^const_name>\undefined\, T)] |> snd |> the_single |> dest_Const |> snd (* TODO: use "Term_Subst.instantiateT" instead? *) fun instantiate_type thy T1 T1' T2 = Same.commit (Envir.subst_type_same (Sign.typ_match thy (T1, T1') Vartab.empty)) T2 handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], []) fun varify_and_instantiate_type ctxt T1 T1' T2 = let val thy = Proof_Context.theory_of ctxt in instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2) end fun free_constructors_of ctxt (Type (s, Ts)) = (case Ctr_Sugar.ctr_sugar_of ctxt s of SOME {ctrs, ...} => map_filter (try dest_Const o Ctr_Sugar.mk_ctr Ts) ctrs | NONE => []) | free_constructors_of _ _ = [] (* Similar to "Nitpick_HOL.bounded_exact_card_of_type". 0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means cardinality 2 or more. The specified default cardinality is returned if the cardinality of the type can't be determined. *) fun tiny_card_of_type ctxt sound assigns default_card T = let val thy = Proof_Context.theory_of ctxt val max = 2 (* 1 would be too small for the "fun" case *) fun aux slack avoid T = if member (op =) avoid T then 0 else case AList.lookup (type_equiv thy) assigns T of SOME k => k | NONE => case T of Type (\<^type_name>\fun\, [T1, T2]) => (case (aux slack avoid T1, aux slack avoid T2) of (k, 1) => if slack andalso k = 0 then 0 else 1 | (0, _) => 0 | (_, 0) => 0 | (k1, k2) => if k1 >= max orelse k2 >= max then max else Int.min (max, Integer.pow k2 k1)) | Type (\<^type_name>\set\, [T']) => aux slack avoid (T' --> \<^typ>\bool\) | \<^typ>\prop\ => 2 | \<^typ>\bool\ => 2 (* optimization *) | \<^typ>\nat\ => 0 (* optimization *) | Type ("Int.int", []) => 0 (* optimization *) | Type (s, _) => (case free_constructors_of ctxt T of constrs as _ :: _ => let val constr_cards = map (Integer.prod o map (aux slack (T :: avoid)) o binder_types o snd) constrs in if exists (curry (op =) 0) constr_cards then 0 else Int.min (max, Integer.sum constr_cards) end | [] => case Typedef.get_info ctxt s of ({abs_type, rep_type, ...}, _) :: _ => if not sound then (* We cheat here by assuming that typedef types are infinite if their underlying type is infinite. This is unsound in general but it's hard to think of a realistic example where this would not be the case. We are also slack with representation types: If a representation type has the form "sigma => tau", we consider it enough to check "sigma" for infiniteness. *) (case varify_and_instantiate_type ctxt (Logic.varifyT_global abs_type) T (Logic.varifyT_global rep_type) |> aux true avoid of 0 => 0 | 1 => 1 | _ => default_card) else default_card | [] => default_card) | TFree _ => (* Very slightly unsound: Type variables are assumed not to be constrained to cardinality 1. (In practice, the user would most likely have used "unit" directly anyway.) *) if not sound andalso default_card = 1 then 2 else default_card | TVar _ => default_card in Int.min (max, aux false [] T) end fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt true [] 0 T <> 0 fun is_type_surely_infinite ctxt sound infinite_Ts T = tiny_card_of_type ctxt sound (map (rpair 0) infinite_Ts) 1 T = 0 (* Simple simplifications to ensure that sort annotations don't leave a trail of spurious "True"s. *) fun s_not \<^Const_>\All T for \Abs (s, T', t')\\ = \<^Const>\Ex T for \Abs (s, T', s_not t')\\ | s_not \<^Const_>\Ex T for \Abs (s, T', t')\\ = \<^Const>\All T for \Abs (s, T', s_not t')\\ | s_not \<^Const_>\implies for t1 t2\ = \<^Const>\conj for t1 \s_not t2\\ | s_not \<^Const_>\conj for t1 t2\ = \<^Const>\disj for \s_not t1\ \s_not t2\\ | s_not \<^Const_>\disj for t1 t2\ = \<^Const>\conj for \s_not t1\ \s_not t2\\ | s_not \<^Const_>\False\ = \<^Const>\True\ | s_not \<^Const_>\True\ = \<^Const>\False\ | s_not \<^Const_>\Not for t\ = t | s_not t = \<^Const>\Not for t\ fun s_conj (\<^Const_>\True\, t2) = t2 | s_conj (t1, \<^Const_>\True\) = t1 | s_conj (\<^Const_>\False\, _) = \<^Const>\False\ | s_conj (_, \<^Const_>\False\) = \<^Const>\False\ | s_conj (t1, t2) = if t1 aconv t2 then t1 else HOLogic.mk_conj (t1, t2) fun s_disj (\<^Const_>\False\, t2) = t2 | s_disj (t1, \<^Const_>\False\) = t1 | s_disj (\<^Const_>\True\, _) = \<^Const>\True\ | s_disj (_, \<^Const_>\True\) = \<^Const>\True\ | s_disj (t1, t2) = if t1 aconv t2 then t1 else HOLogic.mk_disj (t1, t2) fun s_imp (\<^Const_>\True\, t2) = t2 | s_imp (t1, \<^Const_>\False\) = s_not t1 | s_imp (\<^Const_>\False\, _) = \<^Const>\True\ | s_imp (_, \<^Const_>\True\) = \<^Const>\True\ | s_imp p = HOLogic.mk_imp p fun s_iff (\<^Const_>\True\, t2) = t2 | s_iff (t1, \<^Const_>\True\) = t1 | s_iff (\<^Const_>\False\, t2) = s_not t2 | s_iff (t1, \<^Const_>\False\) = s_not t1 | s_iff (t1, t2) = if t1 aconv t2 then \<^Const>\True\ else HOLogic.eq_const HOLogic.boolT $ t1 $ t2 fun close_form t = fold (fn ((s, i), T) => fn t' => Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t'))) (Term.add_vars t []) t val hol_close_form_prefix = "ATP." fun hol_close_form t = fold (fn ((s, i), T) => fn t' => HOLogic.all_const T $ Abs (hol_close_form_prefix ^ s, T, abstract_over (Var ((s, i), T), t'))) (Term.add_vars t []) t fun hol_open_form unprefix (t as Const (\<^const_name>\All\, _) $ Abs (s, T, t')) = (case try unprefix s of SOME s => let val names = Name.make_context (map fst (Term.add_var_names t' [])) val (s, _) = Name.variant s names in hol_open_form unprefix (subst_bound (Var ((s, 0), T), t')) end | NONE => t) | hol_open_form _ t = t fun eta_expand _ t 0 = t | eta_expand Ts (Abs (s, T, t')) n = Abs (s, T, eta_expand (T :: Ts) t' (n - 1)) | eta_expand Ts t n = fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t')) (List.take (binder_types (fastype_of1 (Ts, t)), n)) (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0))) fun cong_extensionalize_term ctxt t = if exists_Const (fn (s, _) => s = \<^const_name>\Not\) t then t |> Skip_Proof.make_thm (Proof_Context.theory_of ctxt) |> Meson.cong_extensionalize_thm ctxt |> Thm.prop_of else t fun is_fun_equality (\<^const_name>\HOL.eq\, Type (_, [Type (\<^type_name>\fun\, _), _])) = true | is_fun_equality _ = false fun abs_extensionalize_term ctxt t = if exists_Const is_fun_equality t then t |> Thm.cterm_of ctxt |> Meson.abs_extensionalize_conv ctxt |> Thm.prop_of |> Logic.dest_equals |> snd else t fun unextensionalize_def t = (case t of \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for lhs rhs\\ => (case strip_comb lhs of (c as Const (_, T), args) => if forall is_Var args andalso not (has_duplicates (op =) args) then \<^Const>\Trueprop for \<^Const>\HOL.eq T for c \fold_rev lambda args rhs\\\ else t | _ => t) | _ => t) (* Converts an elim-rule into an equivalent theorem that does not have the predicate variable. Leaves other theorems unchanged. We simply instantiate the conclusion variable to "False". (Cf. "transform_elim_theorem" in "Meson_Clausify".) *) fun transform_elim_prop t = case Logic.strip_imp_concl t of \<^Const_>\Trueprop for \Var (z, \<^typ>\bool\)\\ => subst_Vars [(z, \<^Const>\False\)] t | Var (z, \<^Type>\prop\) => subst_Vars [(z, \<^prop>\False\)] t | _ => t fun specialize_type thy (s, T) t = let fun subst_for (Const (s', T')) = if s = s' then SOME (Sign.typ_match thy (T', T) Vartab.empty) handle Type.TYPE_MATCH => NONE else NONE | subst_for (t1 $ t2) = (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2) | subst_for (Abs (_, _, t')) = subst_for t' | subst_for _ = NONE in (case subst_for t of SOME subst => Envir.subst_term_types subst t | NONE => raise Type.TYPE_MATCH) end fun strip_subgoal goal i ctxt = let val (t, (frees, params)) = Logic.goal_params (Thm.prop_of goal) i ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free)) val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees) val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees in (rev params, hyp_ts, concl_t) end fun extract_lambda_def dest_head (Const (\<^const_name>\HOL.eq\, _) $ t $ u) = let val (head, args) = strip_comb t in (head |> dest_head |> fst, fold_rev (fn t as Var ((s, _), T) => (fn u => Abs (s, T, abstract_over (t, u))) | _ => raise Fail "expected \"Var\"") args u) end | extract_lambda_def _ _ = raise Fail "malformed lifted lambda" fun short_thm_name ctxt th = let val long = Thm.get_name_hint th val short = Long_Name.base_name long in - if Thm.eq_thm_prop (th, singleton (Attrib.eval_thms ctxt) (Facts.named short, [])) then short - else long + (case try (singleton (Attrib.eval_thms ctxt)) (Facts.named short, []) of + SOME th' => if Thm.eq_thm_prop (th, th') then short else long + | _ => long) end val map_prod = Ctr_Sugar_Util.map_prod (* Compare the length of a list with an integer n while traversing at most n elements of the list. *) fun compare_length_with [] n = if n < 0 then GREATER else if n = 0 then EQUAL else LESS | compare_length_with (_ :: xs) n = if n <= 0 then GREATER else compare_length_with xs (n - 1) end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML @@ -1,167 +1,171 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML Author: Fabian Immler, TU Muenchen Author: Makarius Author: Jasmin Blanchette, TU Muenchen SMT solvers as Sledgehammer provers. *) signature SLEDGEHAMMER_PROVER_SMT = sig type stature = ATP_Problem_Generate.stature type mode = Sledgehammer_Prover.mode type prover = Sledgehammer_Prover.prover val smt_builtins : bool Config.T val smt_triggers : bool Config.T val is_smt_prover : Proof.context -> string -> bool val is_smt_prover_installed : Proof.context -> string -> bool val run_smt_solver : mode -> string -> prover val smts : Proof.context -> string list end; structure Sledgehammer_Prover_SMT : SLEDGEHAMMER_PROVER_SMT = struct open ATP_Util open ATP_Proof open ATP_Problem_Generate open ATP_Proof_Reconstruct open Sledgehammer_Util open Sledgehammer_Proof_Methods open Sledgehammer_Isar open Sledgehammer_ATP_Systems open Sledgehammer_Prover val smt_builtins = Attrib.setup_config_bool \<^binding>\sledgehammer_smt_builtins\ (K true) val smt_triggers = Attrib.setup_config_bool \<^binding>\sledgehammer_smt_triggers\ (K true) val smts = sort_strings o SMT_Config.all_solvers_of val is_smt_prover = member (op =) o smts val is_smt_prover_installed = member (op =) o SMT_Config.available_solvers_of (* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until these are sorted out properly in the SMT module, we must interpret these here. *) val z3_failures = [(101, OutOfResources), (103, MalformedInput), (110, MalformedInput), (112, TimedOut)] val unix_failures = [(134, Crashed), (138, Crashed), (139, Crashed)] val smt_failures = z3_failures @ unix_failures fun failure_of_smt_failure (SMT_Failure.Counterexample genuine) = if genuine then Unprovable else GaveUp | failure_of_smt_failure SMT_Failure.Time_Out = TimedOut | failure_of_smt_failure (SMT_Failure.Abnormal_Termination code) = (case AList.lookup (op =) smt_failures code of SOME failure => failure | NONE => UnknownError ("Abnormal termination with exit code " ^ string_of_int code)) | failure_of_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources | failure_of_smt_failure (SMT_Failure.Other_Failure s) = UnknownError s val is_boring_builtin_typ = not o exists_subtype (member (op =) [\<^typ>\nat\, \<^typ>\int\, HOLogic.realT]) fun smt_filter name ({debug, overlord, max_mono_iters, max_new_mono_instances, type_enc, slices, timeout, ...} : params) state goal i facts options = let val run_timeout = slice_timeout slices timeout val (higher_order, nat_as_int) = (case type_enc of - SOME s => (String.isSubstring "native_higher" s, String.isSubstring "arith" s) - | NONE => (false, false)) + SOME s => (SOME (String.isSubstring "native_higher" s), SOME (String.isSubstring "arith" s)) + | NONE => (NONE, NONE)) fun repair_context ctxt = ctxt |> Context.proof_map (SMT_Config.select_solver name) - |> Config.put SMT_Config.higher_order higher_order - |> Config.put SMT_Config.nat_as_int nat_as_int + |> (case higher_order of + SOME higher_order => Config.put SMT_Config.higher_order higher_order + | NONE => I) + |> (case nat_as_int of + SOME nat_as_int => Config.put SMT_Config.nat_as_int nat_as_int + | NONE => I) |> (if overlord then Config.put SMT_Config.debug_files (overlord_file_location_of_prover name |> (fn (path, name) => path ^ "/" ^ name)) else I) |> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers) |> not (Config.get ctxt smt_builtins) ? (SMT_Builtin.filter_builtins is_boring_builtin_typ #> Config.put SMT_Systems.z3_extensions false) |> repair_monomorph_context max_mono_iters default_max_mono_iters max_new_mono_instances default_max_new_mono_instances val state = Proof.map_context (repair_context) state val ctxt = Proof.context_of state val timer = Timer.startRealTimer () val birth = Timer.checkRealTimer timer val filter_result as {outcome, ...} = SMT_Solver.smt_filter ctxt goal facts i run_timeout options handle exn => if Exn.is_interrupt exn orelse debug then Exn.reraise exn else {outcome = SOME (SMT_Failure.Other_Failure (Runtime.exn_message exn)), fact_ids = NONE, atp_proof = K []} val death = Timer.checkRealTimer timer val run_time = death - birth in {outcome = outcome, filter_result = filter_result, used_from = facts, run_time = run_time} end fun run_smt_solver mode name (params as {debug, verbose, isar_proofs, compress, try0, smt_proofs, minimize, preplay_timeout, ...}) ({state, goal, subgoal, subgoal_count, factss, found_proof, ...} : prover_problem) slice = let val (basic_slice, SMT_Slice options) = slice val facts = facts_of_basic_slice basic_slice factss val ctxt = Proof.context_of state val {outcome, filter_result = {fact_ids, atp_proof, ...}, used_from, run_time} = smt_filter name params state goal subgoal facts options val used_facts = (case fact_ids of NONE => map fst used_from | SOME ids => sort_by fst (map (fst o snd) ids)) val outcome = Option.map failure_of_smt_failure outcome val (preferred_methss, message) = (case outcome of NONE => let val _ = found_proof name; val preferred = if smt_proofs then SMT_Method (if name = "verit" then SMT_Verit "default" else SMT_Z3) else Metis_Method (NONE, NONE); val methss = bunches_of_proof_methods ctxt try0 smt_proofs false liftingN; in ((preferred, methss), fn preplay => let val _ = if verbose then writeln "Generating proof text..." else () fun isar_params () = (verbose, (NONE, NONE), preplay_timeout, compress, try0, minimize, atp_proof (), goal) val one_line_params = (preplay (), proof_banner mode name, subgoal, subgoal_count) val num_chained = length (#facts (Proof.goal state)) in proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained one_line_params end) end | SOME failure => ((Auto_Method (* dummy *), []), fn _ => string_of_atp_failure failure)) in {outcome = outcome, used_facts = used_facts, used_from = used_from, preferred_methss = preferred_methss, run_time = run_time, message = message} end end;