# HG changeset patch # User blanchet # Date 1283410276 -7200 # Node ID e820beaf7d8cc60de099ea5179778f6247940528 # Parent 7cf8beb31e0fb1ab67fe6ee3e7ad3221cbacb6e5# Parent c79e6d536267396bf82b0f033c1cf52bbd4e24ed merged diff -r 7cf8beb31e0f -r e820beaf7d8c doc-src/Nitpick/nitpick.tex --- a/doc-src/Nitpick/nitpick.tex Thu Sep 02 08:40:25 2010 +0200 +++ b/doc-src/Nitpick/nitpick.tex Thu Sep 02 08:51:16 2010 +0200 @@ -1921,7 +1921,7 @@ \end{enum} Default values are indicated in square brackets. Boolean options have a negated -counterpart (e.g., \textit{blocking} vs.\ \textit{no\_blocking}). When setting +counterpart (e.g., \textit{blocking} vs.\ \textit{non\_blocking}). When setting Boolean options, ``= \textit{true}'' may be omitted. \subsection{Mode of Operation} diff -r 7cf8beb31e0f -r e820beaf7d8c doc-src/Sledgehammer/sledgehammer.tex --- a/doc-src/Sledgehammer/sledgehammer.tex Thu Sep 02 08:40:25 2010 +0200 +++ b/doc-src/Sledgehammer/sledgehammer.tex Thu Sep 02 08:51:16 2010 +0200 @@ -333,8 +333,9 @@ \def\opargboolorsmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]} Sledgehammer's options are categorized as follows:\ mode of operation -(\S\ref{mode-of-operation}), problem encoding (\S\ref{problem-encoding}), output -format (\S\ref{output-format}), and timeouts (\S\ref{timeouts}). +(\S\ref{mode-of-operation}), problem encoding (\S\ref{problem-encoding}), +relevance filter (\S\ref{relevance-filter}), output format +(\S\ref{output-format}), and authentication (\S\ref{authentication}). The descriptions below refer to the following syntactic quantities: @@ -349,19 +350,13 @@ \end{enum} Default values are indicated in square brackets. Boolean options have a negated -counterpart (e.g., \textit{debug} vs.\ \textit{no\_debug}). When setting +counterpart (e.g., \textit{blocking} vs.\ \textit{non\_blocking}). When setting Boolean options, ``= \textit{true}'' may be omitted. \subsection{Mode of Operation} \label{mode-of-operation} \begin{enum} -%\optrue{blocking}{non\_blocking} -%Specifies whether the \textbf{sledgehammer} command should operate synchronously. -%The asynchronous (non-blocking) mode lets the user start proving the putative -%theorem while Sledgehammer looks for a counterexample, but it can also be more -%confusing. For technical reasons, automatic runs currently always block. - \opnodefault{atps}{string} Specifies the ATPs (automated theorem provers) to use as a space-separated list (e.g., ``\textit{e}~\textit{spass}''). The following ATPs are supported: @@ -414,6 +409,18 @@ \opnodefault{atp}{string} Alias for \textit{atps}. +\opdefault{timeout}{time}{$\mathbf{60}$ s} +Specifies the maximum amount of time that the ATPs should spend searching for a +proof. For historical reasons, the default value of this option can be +overridden using the option ``Sledgehammer: Time Limit'' from the ``Isabelle'' +menu in Proof General. + +\opfalse{blocking}{non\_blocking} +Specifies whether the \textbf{sledgehammer} command should operate +synchronously. The asynchronous (non-blocking) mode lets the user start proving +the putative theorem manually while Sledgehammer looks for a proof, but it can +also be more confusing. + \opfalse{overlord}{no\_overlord} Specifies whether Sledgehammer should put its temporary files in \texttt{\$ISA\-BELLE\_\allowbreak HOME\_\allowbreak USER}, which is useful for @@ -460,11 +467,11 @@ filter. If the option is set to \textit{smart}, it is set to a value that was empirically found to be appropriate for the ATP. A typical value would be 300. -\opsmartx{theory\_relevant}{theory\_irrelevant} -Specifies whether the theory from which a fact comes should be taken into -consideration by the relevance filter. If the option is set to \textit{smart}, -it is taken to be \textit{true} for SPASS and \textit{false} for the other ATPs; -empirical results suggest that these are the best settings. +%\opsmartx{theory\_relevant}{theory\_irrelevant} +%Specifies whether the theory from which a fact comes should be taken into +%consideration by the relevance filter. If the option is set to \textit{smart}, +%it is taken to be \textit{true} for SPASS and \textit{false} for the other ATPs; +%empirical results suggest that these are the best settings. %\opfalse{defs\_relevant}{defs\_irrelevant} %Specifies whether the definition of constants occurring in the formula to prove @@ -501,15 +508,25 @@ \end{enum} -\subsection{Timeouts} -\label{timeouts} +\subsection{Authentication} +\label{authentication} + +\begin{enum} +\opnodefault{expect}{string} +Specifies the expected outcome, which must be one of the following: \begin{enum} -\opdefault{timeout}{time}{$\mathbf{60}$ s} -Specifies the maximum amount of time that the ATPs should spend looking for a -proof. For historical reasons, the default value of this option can be -overridden using the option ``Sledgehammer: Time Limit'' from the ``Isabelle'' -menu in Proof General. +\item[$\bullet$] \textbf{\textit{some}:} Sledgehammer found a (potentially unsound) proof. +\item[$\bullet$] \textbf{\textit{none}:} Sledgehammer found no proof. +\item[$\bullet$] \textbf{\textit{unknown}:} Sledgehammer encountered some problem. +\end{enum} + +Sledgehammer emits an error (if \textit{blocking} is enabled) or a warning +(otherwise) if the actual outcome differs from the expected outcome. This option +is useful for regression testing. + +\nopagebreak +{\small See also \textit{blocking} (\S\ref{mode-of-operation}).} \end{enum} \let\em=\sl diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/HOL.thy --- a/src/HOL/HOL.thy Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/HOL.thy Thu Sep 02 08:51:16 2010 +0200 @@ -30,6 +30,7 @@ "~~/src/Tools/induct.ML" ("~~/src/Tools/induct_tacs.ML") ("Tools/recfun_codegen.ML") + "Tools/async_manager.ML" "Tools/try.ML" begin diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/IsaMakefile Thu Sep 02 08:51:16 2010 +0200 @@ -267,7 +267,7 @@ $(SRC)/Provers/Arith/combine_numerals.ML \ $(SRC)/Provers/Arith/extract_common_term.ML \ $(SRC)/Tools/Metis/metis.ML \ - Tools/ATP/async_manager.ML \ + Tools/async_manager.ML \ Tools/ATP/atp_problem.ML \ Tools/ATP/atp_systems.ML \ Tools/choice_specification.ML \ @@ -320,10 +320,10 @@ Tools/Sledgehammer/metis_clauses.ML \ Tools/Sledgehammer/metis_tactics.ML \ Tools/Sledgehammer/sledgehammer.ML \ - Tools/Sledgehammer/sledgehammer_fact_filter.ML \ - Tools/Sledgehammer/sledgehammer_fact_minimize.ML \ + Tools/Sledgehammer/sledgehammer_filter.ML \ + Tools/Sledgehammer/sledgehammer_minimize.ML \ Tools/Sledgehammer/sledgehammer_isar.ML \ - Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML \ + Tools/Sledgehammer/sledgehammer_reconstruct.ML \ Tools/Sledgehammer/sledgehammer_translate.ML \ Tools/Sledgehammer/sledgehammer_util.ML \ Tools/SMT/cvc3_solver.ML \ diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Metis_Examples/Abstraction.thy --- a/src/HOL/Metis_Examples/Abstraction.thy Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Metis_Examples/Abstraction.thy Thu Sep 02 08:51:16 2010 +0200 @@ -21,7 +21,7 @@ pset :: "'a set => 'a set" order :: "'a set => ('a * 'a) set" -declare [[ atp_problem_prefix = "Abstraction__Collect_triv" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Collect_triv" ]] lemma (*Collect_triv:*) "a \ {x. P x} ==> P a" proof - assume "a \ {x. P x}" @@ -34,11 +34,11 @@ by (metis mem_Collect_eq) -declare [[ atp_problem_prefix = "Abstraction__Collect_mp" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Collect_mp" ]] lemma "a \ {x. P x --> Q x} ==> a \ {x. P x} ==> a \ {x. Q x}" by (metis Collect_imp_eq ComplD UnE) -declare [[ atp_problem_prefix = "Abstraction__Sigma_triv" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_triv" ]] lemma "(a,b) \ Sigma A B ==> a \ A & b \ B a" proof - assume A1: "(a, b) \ Sigma A B" @@ -51,7 +51,7 @@ lemma Sigma_triv: "(a,b) \ Sigma A B ==> a \ A & b \ B a" by (metis SigmaD1 SigmaD2) -declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect" ]] lemma "(a, b) \ (SIGMA x:A. {y. x = f y}) \ a \ A \ a = f b" (* Metis says this is satisfiable! by (metis CollectD SigmaD1 SigmaD2) @@ -85,7 +85,7 @@ qed -declare [[ atp_problem_prefix = "Abstraction__CLF_eq_in_pp" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_in_pp" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL.{f. f \ pset cl}) ==> f \ pset cl" by (metis Collect_mem_eq SigmaD2) @@ -100,7 +100,7 @@ thus "f \ pset cl" by metis qed -declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Pi" ]] lemma "(cl,f) \ (SIGMA cl: CL. {f. f \ pset cl \ pset cl}) ==> f \ pset cl \ pset cl" @@ -112,7 +112,7 @@ thus "f \ pset cl \ pset cl" by metis qed -declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Int" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Int" ]] lemma "(cl,f) \ (SIGMA cl: CL. {f. f \ pset cl \ cl}) ==> f \ pset cl \ cl" @@ -127,27 +127,27 @@ qed -declare [[ atp_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__Sigma_Collect_Pi_mono" ]] lemma "(cl,f) \ (SIGMA cl: CL. {f. f \ pset cl \ pset cl & monotone f (pset cl) (order cl)}) ==> (f \ pset cl \ pset cl) & (monotone f (pset cl) (order cl))" by auto -declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_subset_Collect_Int" ]] lemma "(cl,f) \ CLF ==> CLF \ (SIGMA cl: CL. {f. f \ pset cl \ cl}) ==> f \ pset cl \ cl" by auto -declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Int" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL. {f. f \ pset cl \ cl}) ==> f \ pset cl \ cl" by auto -declare [[ atp_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_subset_Collect_Pi" ]] lemma "(cl,f) \ CLF ==> CLF \ (SIGMA cl': CL. {f. f \ pset cl' \ pset cl'}) ==> @@ -155,7 +155,7 @@ by fast -declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Pi" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL. {f. f \ pset cl \ pset cl}) ==> @@ -163,46 +163,46 @@ by auto -declare [[ atp_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__CLF_eq_Collect_Pi_mono" ]] lemma "(cl,f) \ CLF ==> CLF = (SIGMA cl: CL. {f. f \ pset cl \ pset cl & monotone f (pset cl) (order cl)}) ==> (f \ pset cl \ pset cl) & (monotone f (pset cl) (order cl))" by auto -declare [[ atp_problem_prefix = "Abstraction__map_eq_zipA" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__map_eq_zipA" ]] lemma "map (%x. (f x, g x)) xs = zip (map f xs) (map g xs)" apply (induct xs) apply (metis map_is_Nil_conv zip.simps(1)) by auto -declare [[ atp_problem_prefix = "Abstraction__map_eq_zipB" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__map_eq_zipB" ]] lemma "map (%w. (w -> w, w \ w)) xs = zip (map (%w. w -> w) xs) (map (%w. w \ w) xs)" apply (induct xs) apply (metis Nil_is_map_conv zip_Nil) by auto -declare [[ atp_problem_prefix = "Abstraction__image_evenA" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__image_evenA" ]] lemma "(%x. Suc(f x)) ` {x. even x} <= A ==> (\x. even x --> Suc(f x) \ A)" by (metis Collect_def image_subset_iff mem_def) -declare [[ atp_problem_prefix = "Abstraction__image_evenB" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__image_evenB" ]] lemma "(%x. f (f x)) ` ((%x. Suc(f x)) ` {x. even x}) <= A ==> (\x. even x --> f (f (Suc(f x))) \ A)"; by (metis Collect_def imageI image_image image_subset_iff mem_def) -declare [[ atp_problem_prefix = "Abstraction__image_curry" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__image_curry" ]] lemma "f \ (%u v. b \ u \ v) ` A ==> \u v. P (b \ u \ v) ==> P(f y)" (*sledgehammer*) by auto -declare [[ atp_problem_prefix = "Abstraction__image_TimesA" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesA" ]] lemma image_TimesA: "(%(x,y). (f x, g y)) ` (A \ B) = (f`A) \ (g`B)" (*sledgehammer*) apply (rule equalityI) (***Even the two inclusions are far too difficult -using [[ atp_problem_prefix = "Abstraction__image_TimesA_simpler"]] +using [[ sledgehammer_problem_prefix = "Abstraction__image_TimesA_simpler"]] ***) apply (rule subsetI) apply (erule imageE) @@ -225,13 +225,13 @@ (*Given the difficulty of the previous problem, these two are probably impossible*) -declare [[ atp_problem_prefix = "Abstraction__image_TimesB" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesB" ]] lemma image_TimesB: "(%(x,y,z). (f x, g y, h z)) ` (A \ B \ C) = (f`A) \ (g`B) \ (h`C)" (*sledgehammer*) by force -declare [[ atp_problem_prefix = "Abstraction__image_TimesC" ]] +declare [[ sledgehammer_problem_prefix = "Abstraction__image_TimesC" ]] lemma image_TimesC: "(%(x,y). (x \ x, y \ y)) ` (A \ B) = ((%x. x \ x) ` A) \ ((%y. y \ y) ` B)" diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Metis_Examples/BT.thy --- a/src/HOL/Metis_Examples/BT.thy Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Metis_Examples/BT.thy Thu Sep 02 08:51:16 2010 +0200 @@ -65,7 +65,7 @@ text {* \medskip BT simplification *} -declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__n_leaves_reflect" ]] lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t" proof (induct t) @@ -81,7 +81,7 @@ by (metis n_leaves.simps(2) nat_add_commute reflect.simps(2)) qed -declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__n_nodes_reflect" ]] lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" proof (induct t) @@ -91,7 +91,7 @@ by (metis add_commute n_nodes.simps(2) reflect.simps(2)) qed -declare [[ atp_problem_prefix = "BT__depth_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__depth_reflect" ]] lemma depth_reflect: "depth (reflect t) = depth t" apply (induct t) @@ -102,14 +102,14 @@ The famous relationship between the numbers of leaves and nodes. *} -declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]] +declare [[ sledgehammer_problem_prefix = "BT__n_leaves_nodes" ]] lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)" apply (induct t) apply (metis n_leaves.simps(1) n_nodes.simps(1)) by auto -declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]] +declare [[ sledgehammer_problem_prefix = "BT__reflect_reflect_ident" ]] lemma reflect_reflect_ident: "reflect (reflect t) = t" apply (induct t) @@ -127,7 +127,7 @@ thus "reflect (reflect (Br a t1 t2)) = Br a t1 t2" by blast qed -declare [[ atp_problem_prefix = "BT__bt_map_ident" ]] +declare [[ sledgehammer_problem_prefix = "BT__bt_map_ident" ]] lemma bt_map_ident: "bt_map (%x. x) = (%y. y)" apply (rule ext) @@ -135,35 +135,35 @@ apply (metis bt_map.simps(1)) by (metis bt_map.simps(2)) -declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] +declare [[ sledgehammer_problem_prefix = "BT__bt_map_appnd" ]] lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)" apply (induct t) apply (metis appnd.simps(1) bt_map.simps(1)) by (metis appnd.simps(2) bt_map.simps(2)) -declare [[ atp_problem_prefix = "BT__bt_map_compose" ]] +declare [[ sledgehammer_problem_prefix = "BT__bt_map_compose" ]] lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)" apply (induct t) apply (metis bt_map.simps(1)) by (metis bt_map.simps(2) o_eq_dest_lhs) -declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__bt_map_reflect" ]] lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)" apply (induct t) apply (metis bt_map.simps(1) reflect.simps(1)) by (metis bt_map.simps(2) reflect.simps(2)) -declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]] +declare [[ sledgehammer_problem_prefix = "BT__preorder_bt_map" ]] lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)" apply (induct t) apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1)) by simp -declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]] +declare [[ sledgehammer_problem_prefix = "BT__inorder_bt_map" ]] lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)" proof (induct t) @@ -178,21 +178,21 @@ case (Br a t1 t2) thus ?case by simp qed -declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]] +declare [[ sledgehammer_problem_prefix = "BT__postorder_bt_map" ]] lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)" apply (induct t) apply (metis Nil_is_map_conv bt_map.simps(1) postorder.simps(1)) by simp -declare [[ atp_problem_prefix = "BT__depth_bt_map" ]] +declare [[ sledgehammer_problem_prefix = "BT__depth_bt_map" ]] lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t" apply (induct t) apply (metis bt_map.simps(1) depth.simps(1)) by simp -declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]] +declare [[ sledgehammer_problem_prefix = "BT__n_leaves_bt_map" ]] lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t" apply (induct t) @@ -213,7 +213,7 @@ using F1 by metis qed -declare [[ atp_problem_prefix = "BT__preorder_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__preorder_reflect" ]] lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)" apply (induct t) @@ -222,7 +222,7 @@ by (metis append.simps(1) append.simps(2) postorder.simps(2) preorder.simps(2) reflect.simps(2) rev.simps(2) rev_append rev_swap) -declare [[ atp_problem_prefix = "BT__inorder_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__inorder_reflect" ]] lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)" apply (induct t) @@ -233,7 +233,7 @@ reflect.simps(2) rev.simps(2) rev_append) *) -declare [[ atp_problem_prefix = "BT__postorder_reflect" ]] +declare [[ sledgehammer_problem_prefix = "BT__postorder_reflect" ]] lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)" apply (induct t) @@ -245,14 +245,14 @@ Analogues of the standard properties of the append function for lists. *} -declare [[ atp_problem_prefix = "BT__appnd_assoc" ]] +declare [[ sledgehammer_problem_prefix = "BT__appnd_assoc" ]] lemma appnd_assoc [simp]: "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)" apply (induct t1) apply (metis appnd.simps(1)) by (metis appnd.simps(2)) -declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]] +declare [[ sledgehammer_problem_prefix = "BT__appnd_Lf2" ]] lemma appnd_Lf2 [simp]: "appnd t Lf = t" apply (induct t) @@ -261,14 +261,14 @@ declare max_add_distrib_left [simp] -declare [[ atp_problem_prefix = "BT__depth_appnd" ]] +declare [[ sledgehammer_problem_prefix = "BT__depth_appnd" ]] lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2" apply (induct t1) apply (metis appnd.simps(1) depth.simps(1) plus_nat.simps(1)) by simp -declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]] +declare [[ sledgehammer_problem_prefix = "BT__n_leaves_appnd" ]] lemma n_leaves_appnd [simp]: "n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2" @@ -277,7 +277,7 @@ semiring_norm(111)) by (simp add: left_distrib) -declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] +declare [[ sledgehammer_problem_prefix = "BT__bt_map_appnd" ]] lemma (*bt_map_appnd:*) "bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)" diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Metis_Examples/BigO.thy --- a/src/HOL/Metis_Examples/BigO.thy Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Metis_Examples/BigO.thy Thu Sep 02 08:51:16 2010 +0200 @@ -15,7 +15,7 @@ definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where "O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}" -declare [[ atp_problem_prefix = "BigO__bigo_pos_const" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_pos_const" ]] lemma bigo_pos_const: "(EX (c::'a::linordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" @@ -124,7 +124,7 @@ {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}" by (auto simp add: bigo_def bigo_pos_const) -declare [[ atp_problem_prefix = "BigO__bigo_elt_subset" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_elt_subset" ]] lemma bigo_elt_subset [intro]: "f : O(g) ==> O(f) <= O(g)" apply (auto simp add: bigo_alt_def) apply (rule_tac x = "ca * c" in exI) @@ -142,12 +142,12 @@ done -declare [[ atp_problem_prefix = "BigO__bigo_refl" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_refl" ]] lemma bigo_refl [intro]: "f : O(f)" apply (auto simp add: bigo_def) by (metis mult_1 order_refl) -declare [[ atp_problem_prefix = "BigO__bigo_zero" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_zero" ]] lemma bigo_zero: "0 : O(g)" apply (auto simp add: bigo_def func_zero) by (metis mult_zero_left order_refl) @@ -230,7 +230,7 @@ apply (auto del: subsetI simp del: bigo_plus_idemp) done -declare [[ atp_problem_prefix = "BigO__bigo_plus_eq" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_plus_eq" ]] lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> O(f + g) = O(f) \ O(g)" apply (rule equalityI) @@ -253,13 +253,13 @@ apply (rule abs_triangle_ineq) apply (metis add_nonneg_nonneg) apply (rule add_mono) -using [[ atp_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] +using [[ sledgehammer_problem_prefix = "BigO__bigo_plus_eq_simpler" ]] (*Found by SPASS; SLOW*) apply (metis le_maxI2 linorder_linear linorder_not_le min_max.sup_absorb1 mult_le_cancel_right order_trans) apply (metis le_maxI2 linorder_not_le mult_le_cancel_right order_trans) done -declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded_alt" ]] lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) @@ -277,14 +277,14 @@ qed text{*So here is the easier (and more natural) problem using transitivity*} -declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (* Version 1: one-line proof *) by (metis abs_ge_self abs_mult order_trans) text{*So here is the easier (and more natural) problem using transitivity*} -declare [[ atp_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded_alt_trans" ]] lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (* Version 2: structured proof *) @@ -299,7 +299,7 @@ apply simp done -declare [[ atp_problem_prefix = "BigO__bigo_bounded2" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_bounded2" ]] lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL x. f x <= lb x + g x ==> f : lb +o O(g)" apply (rule set_minus_imp_plus) @@ -314,13 +314,13 @@ thus "(0\'b) \ f x + - lb x" by (metis not_leE diff_minus less_iff_diff_less_0 less_le_not_le) qed -declare [[ atp_problem_prefix = "BigO__bigo_abs" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_abs" ]] lemma bigo_abs: "(%x. abs(f x)) =o O(f)" apply (unfold bigo_def) apply auto by (metis mult_1 order_refl) -declare [[ atp_problem_prefix = "BigO__bigo_abs2" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_abs2" ]] lemma bigo_abs2: "f =o O(%x. abs(f x))" apply (unfold bigo_def) apply auto @@ -383,7 +383,7 @@ by (simp add: bigo_abs3 [symmetric]) qed -declare [[ atp_problem_prefix = "BigO__bigo_mult" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult" ]] lemma bigo_mult [intro]: "O(f)\O(g) <= O(f * g)" apply (rule subsetI) apply (subst bigo_def) @@ -395,7 +395,7 @@ apply(erule_tac x = x in allE)+ apply(subgoal_tac "c * ca * abs(f x * g x) = (c * abs(f x)) * (ca * abs(g x))") -using [[ atp_problem_prefix = "BigO__bigo_mult_simpler" ]] +using [[ sledgehammer_problem_prefix = "BigO__bigo_mult_simpler" ]] prefer 2 apply (metis mult_assoc mult_left_commute abs_of_pos mult_left_commute @@ -406,14 +406,14 @@ abs_mult has just been done *) by (metis abs_ge_zero mult_mono') -declare [[ atp_problem_prefix = "BigO__bigo_mult2" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult2" ]] lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) (*sledgehammer*); apply (rule_tac x = c in exI) apply clarify apply (drule_tac x = x in spec) -using [[ atp_problem_prefix = "BigO__bigo_mult2_simpler" ]] +using [[ sledgehammer_problem_prefix = "BigO__bigo_mult2_simpler" ]] (*sledgehammer [no luck]*); apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") apply (simp add: mult_ac) @@ -421,11 +421,11 @@ apply (rule abs_ge_zero) done -declare [[ atp_problem_prefix = "BigO__bigo_mult3" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult3" ]] lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" by (metis bigo_mult set_rev_mp set_times_intro) -declare [[ atp_problem_prefix = "BigO__bigo_mult4" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult4" ]] lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) @@ -459,13 +459,13 @@ qed qed -declare [[ atp_problem_prefix = "BigO__bigo_mult6" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult6" ]] lemma bigo_mult6: "ALL x. f x ~= 0 ==> O(f * g) = (f::'a => ('b::linordered_field)) *o O(g)" by (metis bigo_mult2 bigo_mult5 order_antisym) (*proof requires relaxing relevance: 2007-01-25*) -declare [[ atp_problem_prefix = "BigO__bigo_mult7" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult7" ]] declare bigo_mult6 [simp] lemma bigo_mult7: "ALL x. f x ~= 0 ==> O(f * g) <= O(f::'a => ('b::linordered_field)) \ O(g)" @@ -477,7 +477,7 @@ done declare bigo_mult6 [simp del] -declare [[ atp_problem_prefix = "BigO__bigo_mult8" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_mult8" ]] declare bigo_mult7[intro!] lemma bigo_mult8: "ALL x. f x ~= 0 ==> O(f * g) = O(f::'a => ('b::linordered_field)) \ O(g)" @@ -528,7 +528,7 @@ qed qed -declare [[ atp_problem_prefix = "BigO__bigo_plus_absorb" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_plus_absorb" ]] lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff); @@ -555,7 +555,7 @@ lemma bigo_const1: "(%x. c) : O(%x. 1)" by (auto simp add: bigo_def mult_ac) -declare [[ atp_problem_prefix = "BigO__bigo_const2" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_const2" ]] lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" by (metis bigo_const1 bigo_elt_subset); @@ -566,7 +566,7 @@ show "O(\x. c) \ O(\x. 1)" by (metis F1 bigo_elt_subset) qed -declare [[ atp_problem_prefix = "BigO__bigo_const3" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_const3" ]] lemma bigo_const3: "(c::'a::linordered_field) ~= 0 ==> (%x. 1) : O(%x. c)" apply (simp add: bigo_def) by (metis abs_eq_0 left_inverse order_refl) @@ -578,7 +578,7 @@ O(%x. c) = O(%x. 1)" by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) -declare [[ atp_problem_prefix = "BigO__bigo_const_mult1" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult1" ]] lemma bigo_const_mult1: "(%x. c * f x) : O(f)" apply (simp add: bigo_def abs_mult) by (metis le_less) @@ -586,7 +586,7 @@ lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" by (rule bigo_elt_subset, rule bigo_const_mult1) -declare [[ atp_problem_prefix = "BigO__bigo_const_mult3" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult3" ]] lemma bigo_const_mult3: "(c::'a::linordered_field) ~= 0 ==> f : O(%x. c * f x)" apply (simp add: bigo_def) (*sledgehammer [no luck]*) @@ -604,7 +604,7 @@ O(%x. c * f x) = O(f)" by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) -declare [[ atp_problem_prefix = "BigO__bigo_const_mult5" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult5" ]] lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) ~= 0 ==> (%x. c) *o O(f) = O(f)" apply (auto del: subsetI) @@ -624,7 +624,7 @@ done -declare [[ atp_problem_prefix = "BigO__bigo_const_mult6" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_const_mult6" ]] lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times @@ -681,7 +681,7 @@ apply (blast intro: order_trans mult_right_mono abs_ge_self) done -declare [[ atp_problem_prefix = "BigO__bigo_setsum1" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_setsum1" ]] lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> EX c. ALL x y. abs(f x y) <= c * (h x y) ==> (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" @@ -698,7 +698,7 @@ (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" by (rule bigo_setsum1, auto) -declare [[ atp_problem_prefix = "BigO__bigo_setsum3" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_setsum3" ]] lemma bigo_setsum3: "f =o O(h) ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o O(%x. SUM y : A x. abs(l x y * h(k x y)))" @@ -729,7 +729,7 @@ apply (erule set_plus_imp_minus) done -declare [[ atp_problem_prefix = "BigO__bigo_setsum5" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_setsum5" ]] lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> ALL x. 0 <= h x ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o @@ -786,7 +786,7 @@ apply (simp add: func_times) done -declare [[ atp_problem_prefix = "BigO__bigo_fix" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_fix" ]] lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> f =o O(h)" apply (simp add: bigo_alt_def) @@ -849,7 +849,7 @@ apply (erule spec)+ done -declare [[ atp_problem_prefix = "BigO__bigo_lesso1" ]] +declare [[ sledgehammer_problem_prefix = "BigO__bigo_lesso1" ]] lemma bigo_lesso1: "ALL x. f x <= g x ==> f ALL x. 0 <= k x ==> ALL x. k x <= f x ==> k ALL x. 0 <= k x ==> ALL x. g x <= k x ==> f EX C. ALL x. f x <= g x + C * abs(h x)" apply (simp only: lesso_def bigo_alt_def) diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Metis_Examples/Tarski.thy Thu Sep 02 08:51:16 2010 +0200 @@ -409,7 +409,7 @@ (*never proved, 2007-01-22: Tarski__CLF_unnamed_lemma NOT PROVABLE because of the conjunction used in the definition: we don't allow reasoning with rules like conjE, which is essential here.*) -declare [[ atp_problem_prefix = "Tarski__CLF_unnamed_lemma" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_unnamed_lemma" ]] lemma (in CLF) [simp]: "f: pset cl -> pset cl & monotone f (pset cl) (order cl)" apply (insert f_cl) @@ -426,7 +426,7 @@ by (simp add: A_def r_def) (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_CLF_dual" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_CLF_dual" ]] declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp] lemma (in CLF) CLF_dual: "(dual cl, f) \ CLF_set" @@ -454,7 +454,7 @@ subsection {* lemmas for Tarski, lub *} (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_le_flubH" ]] declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] lemma (in CLF) lubH_le_flubH: "H = {x. (x, f x) \ r & x \ A} ==> (lub H cl, f (lub H cl)) \ r" @@ -464,7 +464,7 @@ -- {* @{text "\x:H. (x, f (lub H r)) \ r"} *} apply (rule ballI) (*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_le_flubH_simpler" ]] apply (rule transE) -- {* instantiates @{text "(x, ?z) \ order cl to (x, f x)"}, *} -- {* because of the def of @{text H} *} @@ -482,7 +482,7 @@ CLF.monotone_f[rule del] CL.lub_upper[rule del] (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_flubH_le_lubH" ]] declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] CLF.lubH_le_flubH[simp] @@ -492,7 +492,7 @@ apply (rule_tac t = "H" in ssubst, assumption) apply (rule CollectI) apply (rule conjI) -using [[ atp_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__CLF_flubH_le_lubH_simpler" ]] (*??no longer terminates, with combinators apply (metis CO_refl_on lubH_le_flubH monotone_def monotone_f reflD1 reflD2) *) @@ -506,7 +506,7 @@ (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_is_fixp" ]] (* Single-step version fails. The conjecture clauses refer to local abstraction functions (Frees). *) lemma (in CLF) lubH_is_fixp: @@ -557,7 +557,7 @@ "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" apply (simp add: fix_def) apply (rule conjI) -using [[ atp_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_is_fixp_simpler" ]] apply (metis CO_refl_on lubH_le_flubH refl_onD1) apply (metis antisymE flubH_le_lubH lubH_le_flubH) done @@ -576,7 +576,7 @@ apply (simp_all add: P_def) done -declare [[ atp_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_lubH_least_fixf" ]] lemma (in CLF) lubH_least_fixf: "H = {x. (x, f x) \ r & x \ A} ==> \L. (\y \ fix f A. (y,L) \ r) --> (lub H cl, L) \ r" @@ -584,7 +584,7 @@ done subsection {* Tarski fixpoint theorem 1, first part *} -declare [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_T_thm_1_lub" ]] declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp] lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \ r & x \ A} cl" @@ -592,7 +592,7 @@ apply (rule sym) apply (simp add: P_def) apply (rule lubI) -using [[ atp_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__CLF_T_thm_1_lub_simpler" ]] apply (metis P_def fix_subset) apply (metis Collect_conj_eq Collect_mem_eq Int_commute Int_lower1 lub_in_lattice vimage_def) (*??no longer terminates, with combinators @@ -607,7 +607,7 @@ (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__CLF_glbH_is_fixp" ]] declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp] lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \ r & x \ A} ==> glb H cl \ P" @@ -631,13 +631,13 @@ (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__T_thm_1_glb" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__T_thm_1_glb" ]] (*ALL THEOREMS*) lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \ r & x \ A} cl" (*sledgehammer;*) apply (simp add: glb_dual_lub P_def A_def r_def) apply (rule dualA_iff [THEN subst]) (*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]] (*ALL THEOREMS*) +using [[ sledgehammer_problem_prefix = "Tarski__T_thm_1_glb_simpler" ]] (*ALL THEOREMS*) (*sledgehammer;*) apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff) @@ -646,13 +646,13 @@ subsection {* interval *} -declare [[ atp_problem_prefix = "Tarski__rel_imp_elem" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__rel_imp_elem" ]] declare (in CLF) CO_refl_on[simp] refl_on_def [simp] lemma (in CLF) rel_imp_elem: "(x, y) \ r ==> x \ A" by (metis CO_refl_on refl_onD1) declare (in CLF) CO_refl_on[simp del] refl_on_def [simp del] -declare [[ atp_problem_prefix = "Tarski__interval_subset" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__interval_subset" ]] declare (in CLF) rel_imp_elem[intro] declare interval_def [simp] lemma (in CLF) interval_subset: "[| a \ A; b \ A |] ==> interval r a b \ A" @@ -687,7 +687,7 @@ "[| a \ A; b \ A; S \ interval r a b |]==> S \ A" by (simp add: subset_trans [OF _ interval_subset]) -declare [[ atp_problem_prefix = "Tarski__L_in_interval" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__L_in_interval" ]] (*ALL THEOREMS*) lemma (in CLF) L_in_interval: "[| a \ A; b \ A; S \ interval r a b; S \ {}; isLub S cl L; interval r a b \ {} |] ==> L \ interval r a b" @@ -706,7 +706,7 @@ done (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__G_in_interval" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__G_in_interval" ]] (*ALL THEOREMS*) lemma (in CLF) G_in_interval: "[| a \ A; b \ A; interval r a b \ {}; S \ interval r a b; isGlb S cl G; S \ {} |] ==> G \ interval r a b" @@ -715,7 +715,7 @@ dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub) done -declare [[ atp_problem_prefix = "Tarski__intervalPO" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__intervalPO" ]] (*ALL THEOREMS*) lemma (in CLF) intervalPO: "[| a \ A; b \ A; interval r a b \ {} |] ==> (| pset = interval r a b, order = induced (interval r a b) r |) @@ -783,7 +783,7 @@ lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual] (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__interval_is_sublattice" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__interval_is_sublattice" ]] (*ALL THEOREMS*) lemma (in CLF) interval_is_sublattice: "[| a \ A; b \ A; interval r a b \ {} |] ==> interval r a b <<= cl" @@ -791,7 +791,7 @@ apply (rule sublatticeI) apply (simp add: interval_subset) (*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__interval_is_sublattice_simpler" ]] (*sledgehammer *) apply (rule CompleteLatticeI) apply (simp add: intervalPO) @@ -810,7 +810,7 @@ lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)" by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) -declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) lemma (in CLF) Bot_in_lattice: "Bot cl \ A" (*sledgehammer; *) apply (simp add: Bot_def least_def) @@ -820,12 +820,12 @@ done (*first proved 2007-01-25 after relaxing relevance*) -declare [[ atp_problem_prefix = "Tarski__Top_in_lattice" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__Top_in_lattice" ]] (*ALL THEOREMS*) lemma (in CLF) Top_in_lattice: "Top cl \ A" (*sledgehammer;*) apply (simp add: Top_dual_Bot A_def) (*first proved 2007-01-25 after relaxing relevance*) -using [[ atp_problem_prefix = "Tarski__Top_in_lattice_simpler" ]] (*ALL THEOREMS*) +using [[ sledgehammer_problem_prefix = "Tarski__Top_in_lattice_simpler" ]] (*ALL THEOREMS*) (*sledgehammer*) apply (rule dualA_iff [THEN subst]) apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual) @@ -840,7 +840,7 @@ done (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__Bot_prop" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__Bot_prop" ]] (*ALL THEOREMS*) lemma (in CLF) Bot_prop: "x \ A ==> (Bot cl, x) \ r" (*sledgehammer*) apply (simp add: Bot_dual_Top r_def) @@ -849,12 +849,12 @@ dualA_iff A_def dualPO CL_dualCL CLF_dual) done -declare [[ atp_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__Bot_in_lattice" ]] (*ALL THEOREMS*) lemma (in CLF) Top_intv_not_empty: "x \ A ==> interval r x (Top cl) \ {}" apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE) done -declare [[ atp_problem_prefix = "Tarski__Bot_intv_not_empty" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__Bot_intv_not_empty" ]] (*ALL THEOREMS*) lemma (in CLF) Bot_intv_not_empty: "x \ A ==> interval r (Bot cl) x \ {}" apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem) done @@ -866,7 +866,7 @@ by (simp add: P_def fix_subset po_subset_po) (*first proved 2007-01-25 after relaxing relevance*) -declare [[ atp_problem_prefix = "Tarski__Y_subset_A" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__Y_subset_A" ]] declare (in Tarski) P_def[simp] Y_ss [simp] declare fix_subset [intro] subset_trans [intro] lemma (in Tarski) Y_subset_A: "Y \ A" @@ -882,7 +882,7 @@ by (rule Y_subset_A [THEN lub_in_lattice]) (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__lubY_le_flubY" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__lubY_le_flubY" ]] (*ALL THEOREMS*) lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \ r" (*sledgehammer*) apply (rule lub_least) @@ -891,12 +891,12 @@ apply (rule lubY_in_A) -- {* @{text "Y \ P ==> f x = x"} *} apply (rule ballI) -using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]] (*ALL THEOREMS*) +using [[ sledgehammer_problem_prefix = "Tarski__lubY_le_flubY_simpler" ]] (*ALL THEOREMS*) (*sledgehammer *) apply (rule_tac t = "x" in fix_imp_eq [THEN subst]) apply (erule Y_ss [simplified P_def, THEN subsetD]) -- {* @{text "reduce (f x, f (lub Y cl)) \ r to (x, lub Y cl) \ r"} by monotonicity *} -using [[ atp_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]] (*ALL THEOREMS*) +using [[ sledgehammer_problem_prefix = "Tarski__lubY_le_flubY_simplest" ]] (*ALL THEOREMS*) (*sledgehammer*) apply (rule_tac f = "f" in monotoneE) apply (rule monotone_f) @@ -906,7 +906,7 @@ done (*first proved 2007-01-25 after relaxing relevance*) -declare [[ atp_problem_prefix = "Tarski__intY1_subset" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__intY1_subset" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_subset: "intY1 \ A" (*sledgehammer*) apply (unfold intY1_def) @@ -918,7 +918,7 @@ lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD] (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__intY1_f_closed" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__intY1_f_closed" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_f_closed: "x \ intY1 \ f x \ intY1" (*sledgehammer*) apply (simp add: intY1_def interval_def) @@ -926,7 +926,7 @@ apply (rule transE) apply (rule lubY_le_flubY) -- {* @{text "(f (lub Y cl), f x) \ r"} *} -using [[ atp_problem_prefix = "Tarski__intY1_f_closed_simpler" ]] (*ALL THEOREMS*) +using [[ sledgehammer_problem_prefix = "Tarski__intY1_f_closed_simpler" ]] (*ALL THEOREMS*) (*sledgehammer [has been proved before now...]*) apply (rule_tac f=f in monotoneE) apply (rule monotone_f) @@ -939,13 +939,13 @@ apply (simp add: intY1_def interval_def intY1_elem) done -declare [[ atp_problem_prefix = "Tarski__intY1_func" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__intY1_func" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_func: "(%x: intY1. f x) \ intY1 -> intY1" apply (rule restrict_in_funcset) apply (metis intY1_f_closed restrict_in_funcset) done -declare [[ atp_problem_prefix = "Tarski__intY1_mono" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__intY1_mono" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_mono: "monotone (%x: intY1. f x) intY1 (induced intY1 r)" (*sledgehammer *) @@ -954,7 +954,7 @@ done (*proof requires relaxing relevance: 2007-01-25*) -declare [[ atp_problem_prefix = "Tarski__intY1_is_cl" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__intY1_is_cl" ]] (*ALL THEOREMS*) lemma (in Tarski) intY1_is_cl: "(| pset = intY1, order = induced intY1 r |) \ CompleteLattice" (*sledgehammer*) @@ -967,7 +967,7 @@ done (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__v_in_P" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__v_in_P" ]] (*ALL THEOREMS*) lemma (in Tarski) v_in_P: "v \ P" (*sledgehammer*) apply (unfold P_def) @@ -977,7 +977,7 @@ v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono) done -declare [[ atp_problem_prefix = "Tarski__z_in_interval" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__z_in_interval" ]] (*ALL THEOREMS*) lemma (in Tarski) z_in_interval: "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> z \ intY1" (*sledgehammer *) @@ -991,14 +991,14 @@ apply (simp add: induced_def) done -declare [[ atp_problem_prefix = "Tarski__fz_in_int_rel" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__fz_in_int_rel" ]] (*ALL THEOREMS*) lemma (in Tarski) f'z_in_int_rel: "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> ((%x: intY1. f x) z, z) \ induced intY1 r" apply (metis P_def acc_def fix_imp_eq fix_subset indI reflE restrict_apply subset_eq z_in_interval) done (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__tarski_full_lemma" ]] (*ALL THEOREMS*) +declare [[ sledgehammer_problem_prefix = "Tarski__tarski_full_lemma" ]] (*ALL THEOREMS*) lemma (in Tarski) tarski_full_lemma: "\L. isLub Y (| pset = P, order = induced P r |) L" apply (rule_tac x = "v" in exI) @@ -1028,12 +1028,12 @@ prefer 2 apply (simp add: v_in_P) apply (unfold v_def) (*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__tarski_full_lemma_simpler" ]] (*sledgehammer*) apply (rule indE) apply (rule_tac [2] intY1_subset) (*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]] +using [[ sledgehammer_problem_prefix = "Tarski__tarski_full_lemma_simplest" ]] (*sledgehammer*) apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) apply (simp add: CL_imp_PO intY1_is_cl) @@ -1051,7 +1051,7 @@ (*never proved, 2007-01-22*) -declare [[ atp_problem_prefix = "Tarski__Tarski_full" ]] +declare [[ sledgehammer_problem_prefix = "Tarski__Tarski_full" ]] declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp] Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro] CompleteLatticeI_simp [intro] @@ -1061,7 +1061,7 @@ apply (rule CompleteLatticeI_simp) apply (rule fixf_po, clarify) (*never proved, 2007-01-22*) -using [[ atp_problem_prefix = "Tarski__Tarski_full_simpler" ]] +using [[ sledgehammer_problem_prefix = "Tarski__Tarski_full_simpler" ]] (*sledgehammer*) apply (simp add: P_def A_def r_def) apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro, diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Thu Sep 02 08:51:16 2010 +0200 @@ -290,7 +290,7 @@ | NONE => get_prover (default_atp_name ())) end -type locality = Sledgehammer_Fact_Filter.locality +type locality = Sledgehammer_Filter.locality local @@ -299,20 +299,29 @@ SH_FAIL of int * int | SH_ERROR -fun run_sh prover hard_timeout timeout dir st = +fun run_sh prover_name prover hard_timeout timeout dir st = let - val {context = ctxt, facts, goal} = Proof.goal st + val {context = ctxt, facts = chained_ths, goal} = Proof.goal st val thy = ProofContext.theory_of ctxt + val i = 1 val change_dir = (fn SOME d => Config.put Sledgehammer.dest_dir d | _ => I) val ctxt' = ctxt |> change_dir dir - |> Config.put Sledgehammer.measure_runtime true - val params = Sledgehammer_Isar.default_params thy - [("timeout", Int.toString timeout ^ " s")] + |> Config.put Sledgehammer.measure_run_time true + val params as {full_types, relevance_thresholds, max_relevant, ...} = + Sledgehammer_Isar.default_params thy + [("timeout", Int.toString timeout ^ " s")] + val relevance_override = {add = [], del = [], only = false} + val {default_max_relevant, ...} = ATP_Systems.get_prover thy prover_name + val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal i + val axioms = + Sledgehammer_Filter.relevant_facts ctxt full_types relevance_thresholds + (the_default default_max_relevant max_relevant) + relevance_override chained_ths hyp_ts concl_t val problem = - {subgoal = 1, goal = (ctxt', (facts, goal)), - relevance_override = {add = [], del = [], only = false}, axioms = NONE} + {ctxt = ctxt', goal = goal, subgoal = i, + axioms = map (Sledgehammer_Translate.prepare_axiom ctxt) axioms} val time_limit = (case hard_timeout of NONE => I @@ -352,12 +361,12 @@ val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30) val hard_timeout = AList.lookup (op =) args prover_hard_timeoutK |> Option.map (fst o read_int o explode) - val (msg, result) = run_sh prover hard_timeout timeout dir st + val (msg, result) = run_sh prover_name prover hard_timeout timeout dir st in case result of SH_OK (time_isa, time_atp, names) => let - fun get_thms (name, Sledgehammer_Fact_Filter.Chained) = NONE + fun get_thms (name, Sledgehammer_Filter.Chained) = NONE | get_thms (name, loc) = SOME ((name, loc), thms_of_name (Proof.context_of st) name) in @@ -396,7 +405,7 @@ val params = Sledgehammer_Isar.default_params thy [("atps", prover_name), ("timeout", Int.toString timeout ^ " s")] val minimize = - Sledgehammer_Fact_Minimize.minimize_theorems params 1 (subgoal_count st) + Sledgehammer_Minimize.minimize_theorems params 1 (subgoal_count st) val _ = log separator in case minimize st (these (!named_thms)) of diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML Thu Sep 02 08:51:16 2010 +0200 @@ -5,22 +5,26 @@ structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION = struct +structure SF = Sledgehammer_Filter + val relevance_filter_args = - [("worse_irrel_freq", Sledgehammer_Fact_Filter.worse_irrel_freq), - ("higher_order_irrel_weight", - Sledgehammer_Fact_Filter.higher_order_irrel_weight), - ("abs_rel_weight", Sledgehammer_Fact_Filter.abs_rel_weight), - ("abs_irrel_weight", Sledgehammer_Fact_Filter.abs_irrel_weight), - ("skolem_irrel_weight", Sledgehammer_Fact_Filter.skolem_irrel_weight), - ("intro_bonus", Sledgehammer_Fact_Filter.intro_bonus), - ("elim_bonus", Sledgehammer_Fact_Filter.elim_bonus), - ("simp_bonus", Sledgehammer_Fact_Filter.simp_bonus), - ("local_bonus", Sledgehammer_Fact_Filter.local_bonus), - ("chained_bonus", Sledgehammer_Fact_Filter.chained_bonus), - ("max_imperfect", Sledgehammer_Fact_Filter.max_imperfect), - ("max_imperfect_exp", Sledgehammer_Fact_Filter.max_imperfect_exp), - ("threshold_divisor", Sledgehammer_Fact_Filter.threshold_divisor), - ("ridiculous_threshold", Sledgehammer_Fact_Filter.ridiculous_threshold)] + [("worse_irrel_freq", SF.worse_irrel_freq), + ("higher_order_irrel_weight", SF.higher_order_irrel_weight), + ("abs_rel_weight", SF.abs_rel_weight), + ("abs_irrel_weight", SF.abs_irrel_weight), + ("skolem_irrel_weight", SF.skolem_irrel_weight), + ("theory_const_rel_weight", SF.theory_const_rel_weight), + ("theory_const_irrel_weight", SF.theory_const_irrel_weight), + ("intro_bonus", SF.intro_bonus), + ("elim_bonus", SF.elim_bonus), + ("simp_bonus", SF.simp_bonus), + ("local_bonus", SF.local_bonus), + ("assum_bonus", SF.assum_bonus), + ("chained_bonus", SF.chained_bonus), + ("max_imperfect", SF.max_imperfect), + ("max_imperfect_exp", SF.max_imperfect_exp), + ("threshold_divisor", SF.threshold_divisor), + ("ridiculous_threshold", SF.ridiculous_threshold)] structure Prooftab = Table(type key = int * int val ord = prod_ord int_ord int_ord); @@ -99,15 +103,14 @@ SOME rf => (rf := the (Real.fromString value); false) | NONE => true) - val {relevance_thresholds, full_types, max_relevant, theory_relevant, - ...} = Sledgehammer_Isar.default_params thy args + val {relevance_thresholds, full_types, max_relevant, ...} = + Sledgehammer_Isar.default_params thy args val subgoal = 1 val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal subgoal val facts = - Sledgehammer_Fact_Filter.relevant_facts ctxt full_types + SF.relevant_facts ctxt full_types relevance_thresholds (the_default default_max_relevant max_relevant) - (the_default false theory_relevant) {add = [], del = [], only = false} facts hyp_ts concl_t |> map (fst o fst) val (found_facts, lost_facts) = diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Sledgehammer.thy --- a/src/HOL/Sledgehammer.thy Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Sledgehammer.thy Thu Sep 02 08:51:16 2010 +0200 @@ -10,7 +10,6 @@ theory Sledgehammer imports Plain Hilbert_Choice uses - ("Tools/ATP/async_manager.ML") ("Tools/ATP/atp_problem.ML") ("Tools/ATP/atp_systems.ML") ("~~/src/Tools/Metis/metis.ML") @@ -19,11 +18,11 @@ ("Tools/Sledgehammer/metis_clauses.ML") ("Tools/Sledgehammer/metis_tactics.ML") ("Tools/Sledgehammer/sledgehammer_util.ML") - ("Tools/Sledgehammer/sledgehammer_fact_filter.ML") + ("Tools/Sledgehammer/sledgehammer_filter.ML") ("Tools/Sledgehammer/sledgehammer_translate.ML") - ("Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML") + ("Tools/Sledgehammer/sledgehammer_reconstruct.ML") ("Tools/Sledgehammer/sledgehammer.ML") - ("Tools/Sledgehammer/sledgehammer_fact_minimize.ML") + ("Tools/Sledgehammer/sledgehammer_minimize.ML") ("Tools/Sledgehammer/sledgehammer_isar.ML") begin @@ -89,7 +88,6 @@ apply (simp add: COMBC_def) done -use "Tools/ATP/async_manager.ML" use "Tools/ATP/atp_problem.ML" use "Tools/ATP/atp_systems.ML" setup ATP_Systems.setup @@ -103,12 +101,12 @@ use "Tools/Sledgehammer/metis_tactics.ML" use "Tools/Sledgehammer/sledgehammer_util.ML" -use "Tools/Sledgehammer/sledgehammer_fact_filter.ML" +use "Tools/Sledgehammer/sledgehammer_filter.ML" use "Tools/Sledgehammer/sledgehammer_translate.ML" -use "Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML" +use "Tools/Sledgehammer/sledgehammer_reconstruct.ML" use "Tools/Sledgehammer/sledgehammer.ML" setup Sledgehammer.setup -use "Tools/Sledgehammer/sledgehammer_fact_minimize.ML" +use "Tools/Sledgehammer/sledgehammer_minimize.ML" use "Tools/Sledgehammer/sledgehammer_isar.ML" setup Metis_Tactics.setup diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/ATP/async_manager.ML --- a/src/HOL/Tools/ATP/async_manager.ML Thu Sep 02 08:40:25 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,241 +0,0 @@ -(* Title: HOL/Tools/ATP/async_manager.ML - Author: Fabian Immler, TU Muenchen - Author: Makarius - Author: Jasmin Blanchette, TU Muenchen - -Central manager for asynchronous diagnosis tool threads. -*) - -signature ASYNC_MANAGER = -sig - val launch : - string -> bool -> Time.time -> Time.time -> string -> (unit -> string) - -> unit - val kill_threads : string -> string -> unit - val running_threads : string -> string -> unit - val thread_messages : string -> string -> int option -> unit -end; - -structure Async_Manager : ASYNC_MANAGER = -struct - -(** preferences **) - -val message_store_limit = 20; -val message_display_limit = 5; - - -(** thread management **) - -(* data structures over threads *) - -structure Thread_Heap = Heap -( - type elem = Time.time * Thread.thread; - fun ord ((a, _), (b, _)) = Time.compare (a, b); -); - -fun lookup_thread xs = AList.lookup Thread.equal xs; -fun delete_thread xs = AList.delete Thread.equal xs; -fun update_thread xs = AList.update Thread.equal xs; - - -(* state of thread manager *) - -type state = - {manager: Thread.thread option, - timeout_heap: Thread_Heap.T, - active: (Thread.thread * (string * Time.time * Time.time * string)) list, - canceling: (Thread.thread * (string * Time.time * string)) list, - messages: (string * string) list, - store: (string * string) list} - -fun make_state manager timeout_heap active canceling messages store : state = - {manager = manager, timeout_heap = timeout_heap, active = active, - canceling = canceling, messages = messages, store = store} - -val global_state = Synchronized.var "async_manager" - (make_state NONE Thread_Heap.empty [] [] [] []); - - -(* unregister thread *) - -fun unregister verbose message thread = - Synchronized.change global_state - (fn state as {manager, timeout_heap, active, canceling, messages, store} => - (case lookup_thread active thread of - SOME (tool, birth_time, _, desc) => - let - val active' = delete_thread thread active; - val now = Time.now () - val canceling' = (thread, (tool, now, desc)) :: canceling - val message' = - desc ^ "\n" ^ message ^ - (if verbose then - "\nTotal time: " ^ Int.toString (Time.toMilliseconds - (Time.- (now, birth_time))) ^ " ms." - else - "") - val messages' = (tool, message') :: messages; - val store' = (tool, message') :: - (if length store <= message_store_limit then store - else #1 (chop message_store_limit store)); - in make_state manager timeout_heap active' canceling' messages' store' end - | NONE => state)); - - -(* main manager thread -- only one may exist *) - -val min_wait_time = Time.fromMilliseconds 300; -val max_wait_time = Time.fromSeconds 10; - -fun replace_all bef aft = - let - fun aux seen "" = String.implode (rev seen) - | aux seen s = - if String.isPrefix bef s then - aux seen "" ^ aft ^ aux [] (unprefix bef s) - else - aux (String.sub (s, 0) :: seen) (String.extract (s, 1, NONE)) - in aux [] end - -(* This is a workaround for Proof General's off-by-a-few sendback display bug, - whereby "pr" in "proof" is not highlighted. *) -fun break_into_chunks xs = - maps (space_explode "\000" o replace_all "\n\n" "\000" o snd) xs - -fun print_new_messages () = - case Synchronized.change_result global_state - (fn {manager, timeout_heap, active, canceling, messages, store} => - (messages, make_state manager timeout_heap active canceling [] - store)) of - [] => () - | msgs as (tool, _) :: _ => - let val ss = break_into_chunks msgs in - List.app priority (tool ^ ": " ^ hd ss :: tl ss) - end - -fun check_thread_manager verbose = Synchronized.change global_state - (fn state as {manager, timeout_heap, active, canceling, messages, store} => - if (case manager of SOME thread => Thread.isActive thread | NONE => false) then state - else let val manager = SOME (Toplevel.thread false (fn () => - let - fun time_limit timeout_heap = - (case try Thread_Heap.min timeout_heap of - NONE => Time.+ (Time.now (), max_wait_time) - | SOME (time, _) => time); - - (*action: find threads whose timeout is reached, and interrupt canceling threads*) - fun action {manager, timeout_heap, active, canceling, messages, store} = - let val (timeout_threads, timeout_heap') = - Thread_Heap.upto (Time.now (), Thread.self ()) timeout_heap; - in - if null timeout_threads andalso null canceling then - NONE - else - let - val _ = List.app (Simple_Thread.interrupt o #1) canceling - val canceling' = filter (Thread.isActive o #1) canceling - val state' = make_state manager timeout_heap' active canceling' messages store; - in SOME (map #2 timeout_threads, state') end - end; - in - while Synchronized.change_result global_state - (fn state as {timeout_heap, active, canceling, messages, store, ...} => - if null active andalso null canceling andalso null messages - then (false, make_state NONE timeout_heap active canceling messages store) - else (true, state)) - do - (Synchronized.timed_access global_state (SOME o time_limit o #timeout_heap) action - |> these - |> List.app (unregister verbose "Timed out.\n"); - print_new_messages (); - (*give threads some time to respond to interrupt*) - OS.Process.sleep min_wait_time) - end)) - in make_state manager timeout_heap active canceling messages store end) - - -(* register thread *) - -fun register tool verbose birth_time death_time desc thread = - (Synchronized.change global_state - (fn {manager, timeout_heap, active, canceling, messages, store} => - let - val timeout_heap' = Thread_Heap.insert (death_time, thread) timeout_heap; - val active' = update_thread (thread, (tool, birth_time, death_time, desc)) active; - val state' = make_state manager timeout_heap' active' canceling messages store; - in state' end); - check_thread_manager verbose); - - -fun launch tool verbose birth_time death_time desc f = - (Toplevel.thread true - (fn () => - let - val self = Thread.self () - val _ = register tool verbose birth_time death_time desc self - val message = f () - in unregister verbose message self end); - ()) - - -(** user commands **) - -(* kill threads *) - -fun kill_threads tool workers = Synchronized.change global_state - (fn {manager, timeout_heap, active, canceling, messages, store} => - let - val killing = - map_filter (fn (th, (tool', _, _, desc)) => - if tool' = tool then SOME (th, (tool', Time.now (), desc)) - else NONE) active - val state' = make_state manager timeout_heap [] (killing @ canceling) messages store; - val _ = if null killing then () else priority ("Killed active " ^ workers ^ ".") - in state' end); - - -(* running threads *) - -fun seconds time = string_of_int (Time.toSeconds time) ^ " s" - -fun running_threads tool workers = - let - val {active, canceling, ...} = Synchronized.value global_state; - - val now = Time.now (); - fun running_info (_, (tool', birth_time, death_time, desc)) = - if tool' = tool then - SOME ("Running: " ^ seconds (Time.- (now, birth_time)) ^ " -- " ^ - seconds (Time.- (death_time, now)) ^ " to live:\n" ^ desc) - else - NONE - fun canceling_info (_, (tool', death_time, desc)) = - if tool' = tool then - SOME ("Trying to interrupt thread since " ^ - seconds (Time.- (now, death_time)) ^ ":\n" ^ desc) - else - NONE - val running = - case map_filter running_info active of - [] => ["No " ^ workers ^ " running."] - | ss => "Running " ^ workers ^ ":" :: ss - val interrupting = - case map_filter canceling_info canceling of - [] => [] - | ss => "Trying to interrupt the following " ^ workers ^ ":" :: ss - in priority (space_implode "\n\n" (running @ interrupting)) end - -fun thread_messages tool worker opt_limit = - let - val limit = the_default message_display_limit opt_limit; - val tool_store = Synchronized.value global_state - |> #store |> filter (curry (op =) tool o fst) - val header = - "Recent " ^ worker ^ " messages" ^ - (if length tool_store <= limit then ":" - else " (" ^ string_of_int limit ^ " displayed):"); - in List.app priority (header :: break_into_chunks (#1 (chop limit tool_store))) end - -end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/ATP/atp_systems.ML --- a/src/HOL/Tools/ATP/atp_systems.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Tools/ATP/atp_systems.ML Thu Sep 02 08:51:16 2010 +0200 @@ -20,7 +20,6 @@ proof_delims: (string * string) list, known_failures: (failure * string) list, default_max_relevant: int, - default_theory_relevant: bool, explicit_forall: bool, use_conjecture_for_hypotheses: bool} @@ -53,7 +52,6 @@ proof_delims: (string * string) list, known_failures: (failure * string) list, default_max_relevant: int, - default_theory_relevant: bool, explicit_forall: bool, use_conjecture_for_hypotheses: bool} @@ -160,7 +158,6 @@ (OutOfResources, "SZS status: ResourceOut"), (OutOfResources, "SZS status ResourceOut")], default_max_relevant = 500 (* FUDGE *), - default_theory_relevant = false, explicit_forall = false, use_conjecture_for_hypotheses = true} @@ -171,7 +168,7 @@ counteracting the presence of "hAPP". *) val spass_config : prover_config = {exec = ("ISABELLE_ATP", "scripts/spass"), - required_execs = [("SPASS_HOME", "SPASS")], + required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")], arguments = fn complete => fn timeout => ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 0 timeout)) @@ -187,7 +184,6 @@ (MalformedInput, "Free Variable"), (SpassTooOld, "tptp2dfg")], default_max_relevant = 350 (* FUDGE *), - default_theory_relevant = true, explicit_forall = true, use_conjecture_for_hypotheses = true} @@ -217,7 +213,6 @@ (Unprovable, "Termination reason: Satisfiable"), (VampireTooOld, "not a valid option")], default_max_relevant = 400 (* FUDGE *), - default_theory_relevant = false, explicit_forall = false, use_conjecture_for_hypotheses = true} @@ -252,12 +247,12 @@ fun the_system name versions = case get_system name versions of - NONE => error ("System " ^ quote name ^ " not available at SystemOnTPTP.") - | SOME sys => sys + SOME sys => sys + | NONE => error ("System " ^ quote name ^ " not available at SystemOnTPTP.") fun remote_config system_name system_versions proof_delims known_failures - default_max_relevant default_theory_relevant - use_conjecture_for_hypotheses : prover_config = + default_max_relevant use_conjecture_for_hypotheses + : prover_config = {exec = ("ISABELLE_ATP", "scripts/remote_atp"), required_execs = [], arguments = fn _ => fn timeout => @@ -269,26 +264,21 @@ known_failures @ known_perl_failures @ [(TimedOut, "says Timeout")], default_max_relevant = default_max_relevant, - default_theory_relevant = default_theory_relevant, explicit_forall = true, use_conjecture_for_hypotheses = use_conjecture_for_hypotheses} fun remotify_config system_name system_versions ({proof_delims, known_failures, default_max_relevant, - default_theory_relevant, use_conjecture_for_hypotheses, ...} - : prover_config) : prover_config = + use_conjecture_for_hypotheses, ...} : prover_config) : prover_config = remote_config system_name system_versions proof_delims known_failures - default_max_relevant default_theory_relevant - use_conjecture_for_hypotheses + default_max_relevant use_conjecture_for_hypotheses val remotify_name = prefix "remote_" fun remote_prover name system_name system_versions proof_delims known_failures - default_max_relevant default_theory_relevant - use_conjecture_for_hypotheses = + default_max_relevant use_conjecture_for_hypotheses = (remotify_name name, remote_config system_name system_versions proof_delims known_failures - default_max_relevant default_theory_relevant - use_conjecture_for_hypotheses) + default_max_relevant use_conjecture_for_hypotheses) fun remotify_prover (name, config) system_name system_versions = (remotify_name name, remotify_config system_name system_versions config) @@ -296,10 +286,10 @@ val remote_vampire = remotify_prover vampire "Vampire" ["9.0", "1.0", "0.6"] val remote_sine_e = remote_prover "sine_e" "SInE" [] [] [(Unprovable, "says Unknown")] - 1000 (* FUDGE *) false true + 800 (* FUDGE *) true val remote_snark = remote_prover "snark" "SNARK---" [] [("refutation.", "end_refutation.")] [] - 350 (* FUDGE *) false true + 250 (* FUDGE *) true (* Setup *) diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Nitpick/nitpick_isar.ML --- a/src/HOL/Tools/Nitpick/nitpick_isar.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Thu Sep 02 08:51:16 2010 +0200 @@ -348,7 +348,7 @@ |>> curry (op =) "genuine") in if auto orelse blocking then go () - else (Toplevel.thread true (fn () => (go (); ())); (false, state)) + else (Toplevel.thread true (tap go); (false, state)) end fun nitpick_trans (params, i) = diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/metis_tactics.ML --- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML Thu Sep 02 08:51:16 2010 +0200 @@ -814,7 +814,7 @@ "by (metis X)", to prevent "Subgoal.FOCUS" from freezing the type variables. We don't do it for nonschematic facts "X" because this breaks a few proofs (in the rare and subtle case where a proof relied on extensionality not being - applied) and brings no benefits. *) + applied) and brings few benefits. *) val has_tvar = exists_type (exists_subtype (fn TVar _ => true | _ => false)) o prop_of fun method name mode = diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Thu Sep 02 08:51:16 2010 +0200 @@ -9,11 +9,13 @@ signature SLEDGEHAMMER = sig type failure = ATP_Systems.failure - type locality = Sledgehammer_Fact_Filter.locality - type relevance_override = Sledgehammer_Fact_Filter.relevance_override - type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command + type locality = Sledgehammer_Filter.locality + type relevance_override = Sledgehammer_Filter.relevance_override + type fol_formula = Sledgehammer_Translate.fol_formula + type minimize_command = Sledgehammer_Reconstruct.minimize_command type params = - {debug: bool, + {blocking: bool, + debug: bool, verbose: bool, overlord: bool, atps: string list, @@ -21,15 +23,15 @@ explicit_apply: bool, relevance_thresholds: real * real, max_relevant: int option, - theory_relevant: bool option, isar_proof: bool, isar_shrink_factor: int, - timeout: Time.time} + timeout: Time.time, + expect: string} type problem = - {subgoal: int, - goal: Proof.context * (thm list * thm), - relevance_override: relevance_override, - axioms: ((string * locality) * thm) list option} + {ctxt: Proof.context, + goal: thm, + subgoal: int, + axioms: (term * ((string * locality) * fol_formula) option) list} type prover_result = {outcome: failure option, message: string, @@ -44,7 +46,7 @@ val dest_dir : string Config.T val problem_prefix : string Config.T - val measure_runtime : bool Config.T + val measure_run_time : bool Config.T val kill_atps: unit -> unit val running_atps: unit -> unit val messages: int option -> unit @@ -62,9 +64,9 @@ open ATP_Systems open Metis_Clauses open Sledgehammer_Util -open Sledgehammer_Fact_Filter +open Sledgehammer_Filter open Sledgehammer_Translate -open Sledgehammer_Proof_Reconstruct +open Sledgehammer_Reconstruct (** The Sledgehammer **) @@ -80,7 +82,8 @@ (** problems, results, provers, etc. **) type params = - {debug: bool, + {blocking: bool, + debug: bool, verbose: bool, overlord: bool, atps: string list, @@ -88,16 +91,16 @@ explicit_apply: bool, relevance_thresholds: real * real, max_relevant: int option, - theory_relevant: bool option, isar_proof: bool, isar_shrink_factor: int, - timeout: Time.time} + timeout: Time.time, + expect: string} type problem = - {subgoal: int, - goal: Proof.context * (thm list * thm), - relevance_override: relevance_override, - axioms: ((string * locality) * thm) list option} + {ctxt: Proof.context, + goal: thm, + subgoal: int, + axioms: (term * ((string * locality) * fol_formula) option) list} type prover_result = {outcome: failure option, @@ -114,14 +117,15 @@ (* configuration attributes *) -val (dest_dir, dest_dir_setup) = Attrib.config_string "atp_dest_dir" (K ""); - (*Empty string means create files in Isabelle's temporary files directory.*) +val (dest_dir, dest_dir_setup) = + Attrib.config_string "sledgehammer_dest_dir" (K "") + (* Empty string means create files in Isabelle's temporary files directory. *) val (problem_prefix, problem_prefix_setup) = - Attrib.config_string "atp_problem_prefix" (K "prob"); + Attrib.config_string "sledgehammer_problem_prefix" (K "prob") -val (measure_runtime, measure_runtime_setup) = - Attrib.config_bool "atp_measure_runtime" (K false); +val (measure_run_time, measure_run_time_setup) = + Attrib.config_bool "sledgehammer_measure_run_time" (K false) fun with_path cleanup after f path = Exn.capture f path @@ -172,10 +176,11 @@ |-- Scan.repeat parse_clause_formula_pair val extract_clause_formula_relation = Substring.full #> Substring.position set_ClauseFormulaRelationN - #> snd #> Substring.string #> strip_spaces_except_between_ident_chars - #> explode #> parse_clause_formula_relation #> fst + #> snd #> Substring.position "." #> fst #> Substring.string + #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation + #> fst -(* TODO: move to "Sledgehammer_Proof_Reconstruct.ML" *) +(* TODO: move to "Sledgehammer_Reconstruct" *) fun repair_conjecture_shape_and_theorem_names output conjecture_shape axiom_names = if String.isSubstring set_ClauseFormulaRelationN output then @@ -210,51 +215,31 @@ fun prover_fun atp_name {exec, required_execs, arguments, has_incomplete_mode, proof_delims, - known_failures, default_max_relevant, default_theory_relevant, - explicit_forall, use_conjecture_for_hypotheses} + known_failures, default_max_relevant, explicit_forall, + use_conjecture_for_hypotheses} ({debug, verbose, overlord, full_types, explicit_apply, - relevance_thresholds, max_relevant, theory_relevant, isar_proof, - isar_shrink_factor, timeout, ...} : params) - minimize_command - ({subgoal, goal = (ctxt, (chained_ths, th)), relevance_override, - axioms} : problem) = + max_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params) + minimize_command ({ctxt, goal, subgoal, axioms} : problem) = let - val (_, hyp_ts, concl_t) = strip_subgoal th subgoal - - val print = priority - fun print_v f = () |> verbose ? print o f - fun print_d f = () |> debug ? print o f - - val the_axioms = - case axioms of - SOME axioms => axioms - | NONE => - (relevant_facts ctxt full_types relevance_thresholds - (the_default default_max_relevant max_relevant) - (the_default default_theory_relevant theory_relevant) - relevance_override chained_ths hyp_ts concl_t - |> tap ((fn n => print_v (fn () => - "Selected " ^ string_of_int n ^ " fact" ^ plural_s n ^ - " for " ^ quote atp_name ^ ".")) o length)) - + val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal + val max_relevant = the_default default_max_relevant max_relevant + val axioms = take max_relevant axioms (* path to unique problem file *) val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER" - else Config.get ctxt dest_dir; - val the_problem_prefix = Config.get ctxt problem_prefix; - fun prob_pathname nr = - let - val probfile = - Path.basic ((if overlord then "prob_" ^ atp_name - else the_problem_prefix ^ serial_string ()) - ^ "_" ^ string_of_int nr) - in - if the_dest_dir = "" then File.tmp_path probfile - else if File.exists (Path.explode the_dest_dir) - then Path.append (Path.explode the_dest_dir) probfile - else error ("No such directory: " ^ quote the_dest_dir ^ ".") - end; - - val measure_run_time = verbose orelse Config.get ctxt measure_runtime + else Config.get ctxt dest_dir + val the_problem_prefix = Config.get ctxt problem_prefix + val problem_file_name = + Path.basic ((if overlord then "prob_" ^ atp_name + else the_problem_prefix ^ serial_string ()) + ^ "_" ^ string_of_int subgoal) + val problem_path_name = + if the_dest_dir = "" then + File.tmp_path problem_file_name + else if File.exists (Path.explode the_dest_dir) then + Path.append (Path.explode the_dest_dir) problem_file_name + else + error ("No such directory: " ^ quote the_dest_dir ^ ".") + val measure_run_time = verbose orelse Config.get ctxt measure_run_time val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec) (* write out problem file and call prover *) fun command_line complete timeout probfile = @@ -262,7 +247,7 @@ val core = File.shell_path command ^ " " ^ arguments complete timeout ^ " " ^ File.shell_path probfile in - (if measure_run_time then "TIMEFORMAT='%3U'; { time " ^ core ^ " ; }" + (if measure_run_time then "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }" else "exec " ^ core) ^ " 2>&1" end fun split_time s = @@ -300,14 +285,13 @@ val readable_names = debug andalso overlord val (problem, pool, conjecture_offset, axiom_names) = prepare_problem ctxt readable_names explicit_forall full_types - explicit_apply hyp_ts concl_t the_axioms + explicit_apply hyp_ts concl_t axioms val ss = strings_for_tptp_problem use_conjecture_for_hypotheses problem val _ = File.write_list probfile ss val conjecture_shape = conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1 |> map single - val _ = print_d (fn () => "Running " ^ quote atp_name ^ "...") val timer = Timer.startRealTimer () val result = do_run false (if has_incomplete_mode then @@ -337,7 +321,7 @@ File.write (Path.explode (Path.implode probfile ^ "_proof")) output val ((pool, conjecture_shape, axiom_names), (output, msecs, proof, outcome)) = - with_path cleanup export run_on (prob_pathname subgoal) + with_path cleanup export run_on problem_path_name val (conjecture_shape, axiom_names) = repair_conjecture_shape_and_theorem_names output conjecture_shape axiom_names @@ -346,10 +330,10 @@ NONE => proof_text isar_proof (pool, debug, isar_shrink_factor, ctxt, conjecture_shape) - (full_types, minimize_command, proof, axiom_names, th, subgoal) + (full_types, minimize_command, proof, axiom_names, goal, subgoal) |>> (fn message => message ^ (if verbose then - "\nATP CPU time: " ^ string_of_int msecs ^ " ms." + "\nReal CPU time: " ^ string_of_int msecs ^ " ms." else "")) | SOME failure => (string_for_failure failure, []) @@ -362,49 +346,93 @@ fun get_prover_fun thy name = prover_fun name (get_prover thy name) -fun start_prover_thread (params as {verbose, full_types, timeout, ...}) i n - relevance_override minimize_command proof_state - atp_name = +fun run_prover ctxt (params as {blocking, timeout, expect, ...}) i n + relevance_override minimize_command + (problem as {goal, ...}) (prover, atp_name) = let - val thy = Proof.theory_of proof_state val birth_time = Time.now () val death_time = Time.+ (birth_time, timeout) - val prover = get_prover_fun thy atp_name - val {context = ctxt, facts, goal} = Proof.goal proof_state; val desc = - "ATP " ^ quote atp_name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^ - Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)); + "ATP " ^ quote atp_name ^ " for subgoal " ^ string_of_int i ^ ":" ^ + (if blocking then + "" + else + "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))) + fun run () = + let + val (outcome_code, message) = + prover_fun atp_name prover params (minimize_command atp_name) problem + |> (fn {outcome, message, ...} => + (if is_some outcome then "none" else "some", message)) + handle ERROR message => ("unknown", "Error: " ^ message ^ "\n") + | exn => ("unknown", "Internal error:\n" ^ + ML_Compiler.exn_message exn ^ "\n") + in + if expect = "" orelse outcome_code = expect then + () + else if blocking then + error ("Unexpected outcome: " ^ quote outcome_code ^ ".") + else + warning ("Unexpected outcome: " ^ quote outcome_code ^ "."); + message + end in - Async_Manager.launch das_Tool verbose birth_time death_time desc - (fn () => - let - val problem = - {subgoal = i, goal = (ctxt, (facts, goal)), - relevance_override = relevance_override, axioms = NONE} - in - prover params (minimize_command atp_name) problem |> #message - handle ERROR message => "Error: " ^ message ^ "\n" - | exn => "Internal error: \n" ^ ML_Compiler.exn_message exn ^ - "\n" - end) + if blocking then priority (desc ^ "\n" ^ TimeLimit.timeLimit timeout run ()) + else Async_Manager.launch das_Tool birth_time death_time desc run end -fun run_sledgehammer {atps = [], ...} _ _ _ _ = error "No ATP is set." - | run_sledgehammer (params as {atps, ...}) i relevance_override - minimize_command state = - case subgoal_count state of +fun run_sledgehammer (params as {blocking, verbose, atps, full_types, + relevance_thresholds, max_relevant, ...}) + i relevance_override minimize_command state = + if null atps then error "No ATP is set." + else case subgoal_count state of 0 => priority "No subgoal!" | n => let - val _ = kill_atps () + val timer = Timer.startRealTimer () + val _ = () |> not blocking ? kill_atps val _ = priority "Sledgehammering..." - val _ = app (start_prover_thread params i n relevance_override - minimize_command state) atps - in () end + fun run () = + let + val {context = ctxt, facts = chained_ths, goal} = Proof.goal state + val thy = Proof.theory_of state + val (_, hyp_ts, concl_t) = strip_subgoal goal i + val provers = map (`(get_prover thy)) atps + val max_max_relevant = + case max_relevant of + SOME n => n + | NONE => fold (Integer.max o #default_max_relevant o fst) + provers 0 + val axioms = + relevant_facts ctxt full_types relevance_thresholds + max_max_relevant relevance_override chained_ths + hyp_ts concl_t + val problem = + {ctxt = ctxt, goal = goal, subgoal = i, + axioms = map (prepare_axiom ctxt) axioms} + val num_axioms = length axioms + val _ = if verbose then + priority ("Selected " ^ string_of_int num_axioms ^ + " fact" ^ plural_s num_axioms ^ ".") + else + () + val _ = + (if blocking then Par_List.map else map) + (run_prover ctxt params i n relevance_override + minimize_command problem) provers + in + if verbose andalso blocking then + priority ("Total time: " ^ + signed_string_of_int (Time.toMilliseconds + (Timer.checkRealTimer timer)) ^ " ms.") + else + () + end + in if blocking then run () else Toplevel.thread true (tap run) |> K () end val setup = dest_dir_setup #> problem_prefix_setup - #> measure_runtime_setup + #> measure_run_time_setup end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Thu Sep 02 08:40:25 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,800 +0,0 @@ -(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML - Author: Jia Meng, Cambridge University Computer Laboratory and NICTA - Author: Jasmin Blanchette, TU Muenchen -*) - -signature SLEDGEHAMMER_FACT_FILTER = -sig - datatype locality = General | Intro | Elim | Simp | Local | Chained - - type relevance_override = - {add: Facts.ref list, - del: Facts.ref list, - only: bool} - - val trace : bool Unsynchronized.ref - val worse_irrel_freq : real Unsynchronized.ref - val higher_order_irrel_weight : real Unsynchronized.ref - val abs_rel_weight : real Unsynchronized.ref - val abs_irrel_weight : real Unsynchronized.ref - val skolem_irrel_weight : real Unsynchronized.ref - val intro_bonus : real Unsynchronized.ref - val elim_bonus : real Unsynchronized.ref - val simp_bonus : real Unsynchronized.ref - val local_bonus : real Unsynchronized.ref - val chained_bonus : real Unsynchronized.ref - val max_imperfect : real Unsynchronized.ref - val max_imperfect_exp : real Unsynchronized.ref - val threshold_divisor : real Unsynchronized.ref - val ridiculous_threshold : real Unsynchronized.ref - val name_thm_pairs_from_ref : - Proof.context -> unit Symtab.table -> thm list -> Facts.ref - -> ((string * locality) * thm) list - val relevant_facts : - Proof.context -> bool -> real * real -> int -> bool -> relevance_override - -> thm list -> term list -> term -> ((string * locality) * thm) list -end; - -structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER = -struct - -open Sledgehammer_Util - -val trace = Unsynchronized.ref false -fun trace_msg msg = if !trace then tracing (msg ()) else () - -(* experimental feature *) -val term_patterns = false - -val respect_no_atp = true - -datatype locality = General | Intro | Elim | Simp | Local | Chained - -type relevance_override = - {add: Facts.ref list, - del: Facts.ref list, - only: bool} - -val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator - -fun repair_name reserved multi j name = - (name |> Symtab.defined reserved name ? quote) ^ - (if multi then "(" ^ Int.toString j ^ ")" else "") - -fun name_thm_pairs_from_ref ctxt reserved chained_ths xref = - let - val ths = ProofContext.get_fact ctxt xref - val name = Facts.string_of_ref xref - val multi = length ths > 1 - in - (ths, (1, [])) - |-> fold (fn th => fn (j, rest) => - (j + 1, ((repair_name reserved multi j name, - if member Thm.eq_thm chained_ths th then Chained - else General), th) :: rest)) - |> snd - end - -(***************************************************************) -(* Relevance Filtering *) -(***************************************************************) - -(*** constants with types ***) - -fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) = - order_of_type T1 (* cheat: pretend sets are first-order *) - | order_of_type (Type (@{type_name fun}, [T1, T2])) = - Int.max (order_of_type T1 + 1, order_of_type T2) - | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0 - | order_of_type _ = 0 - -(* An abstraction of Isabelle types and first-order terms *) -datatype pattern = PVar | PApp of string * pattern list -datatype ptype = PType of int * pattern list - -fun string_for_pattern PVar = "_" - | string_for_pattern (PApp (s, ps)) = - if null ps then s else s ^ string_for_patterns ps -and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")" -fun string_for_ptype (PType (_, ps)) = string_for_patterns ps - -(*Is the second type an instance of the first one?*) -fun match_pattern (PVar, _) = true - | match_pattern (PApp _, PVar) = false - | match_pattern (PApp (s, ps), PApp (t, qs)) = - s = t andalso match_patterns (ps, qs) -and match_patterns (_, []) = true - | match_patterns ([], _) = false - | match_patterns (p :: ps, q :: qs) = - match_pattern (p, q) andalso match_patterns (ps, qs) -fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs) - -(* Is there a unifiable constant? *) -fun pconst_mem f consts (s, ps) = - exists (curry (match_ptype o f) ps) - (map snd (filter (curry (op =) s o fst) consts)) -fun pconst_hyper_mem f const_tab (s, ps) = - exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s)) - -fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts) - | pattern_for_type (TFree (s, _)) = PApp (s, []) - | pattern_for_type (TVar _) = PVar - -fun pterm thy t = - case strip_comb t of - (Const x, ts) => PApp (pconst thy true x ts) - | (Free x, ts) => PApp (pconst thy false x ts) - | (Var x, []) => PVar - | _ => PApp ("?", []) (* equivalence class of higher-order constructs *) -(* Pairs a constant with the list of its type instantiations. *) -and ptype thy const x ts = - (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x)) - else []) @ - (if term_patterns then map (pterm thy) ts else []) -and pconst thy const (s, T) ts = (s, ptype thy const (s, T) ts) -and rich_ptype thy const (s, T) ts = - PType (order_of_type T, ptype thy const (s, T) ts) -and rich_pconst thy const (s, T) ts = (s, rich_ptype thy const (s, T) ts) - -fun string_for_hyper_pconst (s, ps) = - s ^ "{" ^ commas (map string_for_ptype ps) ^ "}" - -val abs_name = "Sledgehammer.abs" -val skolem_prefix = "Sledgehammer.sko" - -(* These are typically simplified away by "Meson.presimplify". Equality is - handled specially via "fequal". *) -val boring_consts = - [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let}, - @{const_name HOL.eq}] - -(* Add a pconstant to the table, but a [] entry means a standard - connective, which we ignore.*) -fun add_pconst_to_table also_skolem (c, p) = - if member (op =) boring_consts c orelse - (not also_skolem andalso String.isPrefix skolem_prefix c) then - I - else - Symtab.map_default (c, [p]) (insert (op =) p) - -fun is_formula_type T = (T = HOLogic.boolT orelse T = propT) - -fun pconsts_in_terms thy also_skolems pos ts = - let - val flip = Option.map not - (* We include free variables, as well as constants, to handle locales. For - each quantifiers that must necessarily be skolemized by the ATP, we - introduce a fresh constant to simulate the effect of Skolemization. *) - fun do_const const (s, T) ts = - add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts) - #> fold do_term ts - and do_term t = - case strip_comb t of - (Const x, ts) => do_const true x ts - | (Free x, ts) => do_const false x ts - | (Abs (_, T, t'), ts) => - (null ts - ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, []))) - #> fold do_term (t' :: ts) - | (_, ts) => fold do_term ts - fun do_quantifier will_surely_be_skolemized abs_T body_t = - do_formula pos body_t - #> (if also_skolems andalso will_surely_be_skolemized then - add_pconst_to_table true - (gensym skolem_prefix, PType (order_of_type abs_T, [])) - else - I) - and do_term_or_formula T = - if is_formula_type T then do_formula NONE else do_term - and do_formula pos t = - case t of - Const (@{const_name all}, _) $ Abs (_, T, t') => - do_quantifier (pos = SOME false) T t' - | @{const "==>"} $ t1 $ t2 => - do_formula (flip pos) t1 #> do_formula pos t2 - | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 => - fold (do_term_or_formula T) [t1, t2] - | @{const Trueprop} $ t1 => do_formula pos t1 - | @{const Not} $ t1 => do_formula (flip pos) t1 - | Const (@{const_name All}, _) $ Abs (_, T, t') => - do_quantifier (pos = SOME false) T t' - | Const (@{const_name Ex}, _) $ Abs (_, T, t') => - do_quantifier (pos = SOME true) T t' - | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2] - | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2] - | @{const HOL.implies} $ t1 $ t2 => - do_formula (flip pos) t1 #> do_formula pos t2 - | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 => - fold (do_term_or_formula T) [t1, t2] - | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])])) - $ t1 $ t2 $ t3 => - do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3] - | Const (@{const_name Ex1}, _) $ Abs (_, T, t') => - do_quantifier (is_some pos) T t' - | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') => - do_quantifier (pos = SOME false) T - (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t')) - | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') => - do_quantifier (pos = SOME true) T - (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t')) - | (t0 as Const (_, @{typ bool})) $ t1 => - do_term t0 #> do_formula pos t1 (* theory constant *) - | _ => do_term t - in Symtab.empty |> fold (do_formula pos) ts end - -(*Inserts a dummy "constant" referring to the theory name, so that relevance - takes the given theory into account.*) -fun theory_const_prop_of theory_relevant th = - if theory_relevant then - let - val name = Context.theory_name (theory_of_thm th) - val t = Const (name ^ ". 1", @{typ bool}) - in t $ prop_of th end - else - prop_of th - -(**** Constant / Type Frequencies ****) - -(* A two-dimensional symbol table counts frequencies of constants. It's keyed - first by constant name and second by its list of type instantiations. For the - latter, we need a linear ordering on "pattern list". *) - -fun pattern_ord p = - case p of - (PVar, PVar) => EQUAL - | (PVar, PApp _) => LESS - | (PApp _, PVar) => GREATER - | (PApp q1, PApp q2) => - prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2) -fun ptype_ord (PType p, PType q) = - prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q) - -structure PType_Tab = Table(type key = ptype val ord = ptype_ord) - -fun count_axiom_consts theory_relevant thy = - let - fun do_const const (s, T) ts = - (* Two-dimensional table update. Constant maps to types maps to count. *) - PType_Tab.map_default (rich_ptype thy const (s, T) ts, 0) (Integer.add 1) - |> Symtab.map_default (s, PType_Tab.empty) - #> fold do_term ts - and do_term t = - case strip_comb t of - (Const x, ts) => do_const true x ts - | (Free x, ts) => do_const false x ts - | (Abs (_, _, t'), ts) => fold do_term (t' :: ts) - | (_, ts) => fold do_term ts - in do_term o theory_const_prop_of theory_relevant o snd end - - -(**** Actual Filtering Code ****) - -fun pow_int x 0 = 1.0 - | pow_int x 1 = x - | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x - -(*The frequency of a constant is the sum of those of all instances of its type.*) -fun pconst_freq match const_tab (c, ps) = - PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m) - (the (Symtab.lookup const_tab c)) 0 - - -(* A surprising number of theorems contain only a few significant constants. - These include all induction rules, and other general theorems. *) - -(* "log" seems best in practice. A constant function of one ignores the constant - frequencies. Rare constants give more points if they are relevant than less - rare ones. *) -fun rel_weight_for order freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0) - -(* FUDGE *) -val worse_irrel_freq = Unsynchronized.ref 100.0 -val higher_order_irrel_weight = Unsynchronized.ref 1.05 - -(* Irrelevant constants are treated differently. We associate lower penalties to - very rare constants and very common ones -- the former because they can't - lead to the inclusion of too many new facts, and the latter because they are - so common as to be of little interest. *) -fun irrel_weight_for order freq = - let val (k, x) = !worse_irrel_freq |> `Real.ceil in - (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x - else rel_weight_for order freq / rel_weight_for order k) - * pow_int (!higher_order_irrel_weight) (order - 1) - end - -(* FUDGE *) -val abs_rel_weight = Unsynchronized.ref 0.5 -val abs_irrel_weight = Unsynchronized.ref 2.0 -val skolem_irrel_weight = Unsynchronized.ref 0.75 - -(* Computes a constant's weight, as determined by its frequency. *) -fun generic_pconst_weight abs_weight skolem_weight weight_for f const_tab - (c as (s, PType (m, _))) = - if s = abs_name then abs_weight - else if String.isPrefix skolem_prefix s then skolem_weight - else weight_for m (pconst_freq (match_ptype o f) const_tab c) - -fun rel_pconst_weight const_tab = - generic_pconst_weight (!abs_rel_weight) 0.0 rel_weight_for I const_tab -fun irrel_pconst_weight const_tab = - generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight) - irrel_weight_for swap const_tab - -(* FUDGE *) -val intro_bonus = Unsynchronized.ref 0.15 -val elim_bonus = Unsynchronized.ref 0.15 -val simp_bonus = Unsynchronized.ref 0.15 -val local_bonus = Unsynchronized.ref 0.55 -val chained_bonus = Unsynchronized.ref 1.5 - -fun locality_bonus General = 0.0 - | locality_bonus Intro = !intro_bonus - | locality_bonus Elim = !elim_bonus - | locality_bonus Simp = !simp_bonus - | locality_bonus Local = !local_bonus - | locality_bonus Chained = !chained_bonus - -fun axiom_weight loc const_tab relevant_consts axiom_consts = - case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts) - ||> filter_out (pconst_hyper_mem swap relevant_consts) of - ([], _) => 0.0 - | (rel, irrel) => - let - val irrel = irrel |> filter_out (pconst_mem swap rel) - val rel_weight = - 0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel - val irrel_weight = - ~ (locality_bonus loc) - |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel - val res = rel_weight / (rel_weight + irrel_weight) - in if Real.isFinite res then res else 0.0 end - -(* FIXME: experiment -fun debug_axiom_weight loc const_tab relevant_consts axiom_consts = - case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts) - ||> filter_out (pconst_hyper_mem swap relevant_consts) of - ([], _) => 0.0 - | (rel, irrel) => - let - val irrel = irrel |> filter_out (pconst_mem swap rel) - val rels_weight = - 0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel - val irrels_weight = - ~ (locality_bonus loc) - |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel -val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel)) -val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel)) - val res = rels_weight / (rels_weight + irrels_weight) - in if Real.isFinite res then res else 0.0 end -*) - -fun pconsts_in_axiom thy t = - Symtab.fold (fn (s, pss) => fold (cons o pair s) pss) - (pconsts_in_terms thy true (SOME true) [t]) [] -fun pair_consts_axiom theory_relevant thy axiom = - case axiom |> snd |> theory_const_prop_of theory_relevant - |> pconsts_in_axiom thy of - [] => NONE - | consts => SOME ((axiom, consts), NONE) - -type annotated_thm = - (((unit -> string) * locality) * thm) * (string * ptype) list - -(* FUDGE *) -val max_imperfect = Unsynchronized.ref 11.5 -val max_imperfect_exp = Unsynchronized.ref 1.0 - -fun take_most_relevant max_relevant remaining_max - (candidates : (annotated_thm * real) list) = - let - val max_imperfect = - Real.ceil (Math.pow (!max_imperfect, - Math.pow (Real.fromInt remaining_max - / Real.fromInt max_relevant, !max_imperfect_exp))) - val (perfect, imperfect) = - candidates |> sort (Real.compare o swap o pairself snd) - |> take_prefix (fn (_, w) => w > 0.99999) - val ((accepts, more_rejects), rejects) = - chop max_imperfect imperfect |>> append perfect |>> chop remaining_max - in - trace_msg (fn () => - "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^ - Int.toString (length candidates) ^ "): " ^ - (accepts |> map (fn ((((name, _), _), _), weight) => - name () ^ " [" ^ Real.toString weight ^ "]") - |> commas)); - (accepts, more_rejects @ rejects) - end - -fun if_empty_replace_with_locality thy axioms loc tab = - if Symtab.is_empty tab then - pconsts_in_terms thy false (SOME false) - (map_filter (fn ((_, loc'), th) => - if loc' = loc then SOME (prop_of th) else NONE) axioms) - else - tab - -(* FUDGE *) -val threshold_divisor = Unsynchronized.ref 2.0 -val ridiculous_threshold = Unsynchronized.ref 0.1 - -fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant - ({add, del, ...} : relevance_override) axioms goal_ts = - let - val thy = ProofContext.theory_of ctxt - val const_tab = - fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty - val goal_const_tab = - pconsts_in_terms thy false (SOME false) goal_ts - |> fold (if_empty_replace_with_locality thy axioms) [Chained, Local] - val add_thms = maps (ProofContext.get_fact ctxt) add - val del_thms = maps (ProofContext.get_fact ctxt) del - fun iter j remaining_max threshold rel_const_tab hopeless hopeful = - let - fun game_over rejects = - (* Add "add:" facts. *) - if null add_thms then - [] - else - map_filter (fn ((p as (_, th), _), _) => - if member Thm.eq_thm add_thms th then SOME p - else NONE) rejects - fun relevant [] rejects [] = - (* Nothing has been added this iteration. *) - if j = 0 andalso threshold >= !ridiculous_threshold then - (* First iteration? Try again. *) - iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab - hopeless hopeful - else - game_over (rejects @ hopeless) - | relevant candidates rejects [] = - let - val (accepts, more_rejects) = - take_most_relevant max_relevant remaining_max candidates - val rel_const_tab' = - rel_const_tab - |> fold (add_pconst_to_table false) (maps (snd o fst) accepts) - fun is_dirty (c, _) = - Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c - val (hopeful_rejects, hopeless_rejects) = - (rejects @ hopeless, ([], [])) - |-> fold (fn (ax as (_, consts), old_weight) => - if exists is_dirty consts then - apfst (cons (ax, NONE)) - else - apsnd (cons (ax, old_weight))) - |>> append (more_rejects - |> map (fn (ax as (_, consts), old_weight) => - (ax, if exists is_dirty consts then NONE - else SOME old_weight))) - val threshold = - 1.0 - (1.0 - threshold) - * Math.pow (decay, Real.fromInt (length accepts)) - val remaining_max = remaining_max - length accepts - in - trace_msg (fn () => "New or updated constants: " ^ - commas (rel_const_tab' |> Symtab.dest - |> subtract (op =) (rel_const_tab |> Symtab.dest) - |> map string_for_hyper_pconst)); - map (fst o fst) accepts @ - (if remaining_max = 0 then - game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects) - else - iter (j + 1) remaining_max threshold rel_const_tab' - hopeless_rejects hopeful_rejects) - end - | relevant candidates rejects - (((ax as (((_, loc), th), axiom_consts)), cached_weight) - :: hopeful) = - let - val weight = - case cached_weight of - SOME w => w - | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts -(* FIXME: experiment -val name = fst (fst (fst ax)) () -val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then -tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts)) -else -() -*) - in - if weight >= threshold then - relevant ((ax, weight) :: candidates) rejects hopeful - else - relevant candidates ((ax, weight) :: rejects) hopeful - end - in - trace_msg (fn () => - "ITERATION " ^ string_of_int j ^ ": current threshold: " ^ - Real.toString threshold ^ ", constants: " ^ - commas (rel_const_tab |> Symtab.dest - |> filter (curry (op <>) [] o snd) - |> map string_for_hyper_pconst)); - relevant [] [] hopeful - end - in - axioms |> filter_out (member Thm.eq_thm del_thms o snd) - |> map_filter (pair_consts_axiom theory_relevant thy) - |> iter 0 max_relevant threshold0 goal_const_tab [] - |> tap (fn res => trace_msg (fn () => - "Total relevant: " ^ Int.toString (length res))) - end - - -(***************************************************************) -(* Retrieving and filtering lemmas *) -(***************************************************************) - -(*** retrieve lemmas and filter them ***) - -(*Reject theorems with names like "List.filter.filter_list_def" or - "Accessible_Part.acc.defs", as these are definitions arising from packages.*) -fun is_package_def a = - let val names = Long_Name.explode a - in - length names > 2 andalso - not (hd names = "local") andalso - String.isSuffix "_def" a orelse String.isSuffix "_defs" a - end; - -fun mk_fact_table f xs = - fold (Termtab.update o `(prop_of o f)) xs Termtab.empty -fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) [] - -(* FIXME: put other record thms here, or declare as "no_atp" *) -val multi_base_blacklist = - ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits", - "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy", - "case_cong", "weak_case_cong"] - |> map (prefix ".") - -val max_lambda_nesting = 3 - -fun term_has_too_many_lambdas max (t1 $ t2) = - exists (term_has_too_many_lambdas max) [t1, t2] - | term_has_too_many_lambdas max (Abs (_, _, t)) = - max = 0 orelse term_has_too_many_lambdas (max - 1) t - | term_has_too_many_lambdas _ _ = false - -(* Don't count nested lambdas at the level of formulas, since they are - quantifiers. *) -fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) = - formula_has_too_many_lambdas (T :: Ts) t - | formula_has_too_many_lambdas Ts t = - if is_formula_type (fastype_of1 (Ts, t)) then - exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t)) - else - term_has_too_many_lambdas max_lambda_nesting t - -(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31) - was 11. *) -val max_apply_depth = 15 - -fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1) - | apply_depth (Abs (_, _, t)) = apply_depth t - | apply_depth _ = 0 - -fun is_formula_too_complex t = - apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t - -val exists_sledgehammer_const = - exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s) - -(* FIXME: make more reliable *) -val exists_low_level_class_const = - exists_Const (fn (s, _) => - String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s) - -fun is_metastrange_theorem th = - case head_of (concl_of th) of - Const (a, _) => (a <> @{const_name Trueprop} andalso - a <> @{const_name "=="}) - | _ => false - -fun is_that_fact th = - String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th) - andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN - | _ => false) (prop_of th) - -val type_has_top_sort = - exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false) - -(**** Predicates to detect unwanted facts (prolific or likely to cause - unsoundness) ****) - -(* Too general means, positive equality literal with a variable X as one - operand, when X does not occur properly in the other operand. This rules out - clearly inconsistent facts such as X = a | X = b, though it by no means - guarantees soundness. *) - -(* Unwanted equalities are those between a (bound or schematic) variable that - does not properly occur in the second operand. *) -val is_exhaustive_finite = - let - fun is_bad_equal (Var z) t = - not (exists_subterm (fn Var z' => z = z' | _ => false) t) - | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j)) - | is_bad_equal _ _ = false - fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1 - fun do_formula pos t = - case (pos, t) of - (_, @{const Trueprop} $ t1) => do_formula pos t1 - | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) => - do_formula pos t' - | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) => - do_formula pos t' - | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) => - do_formula pos t' - | (_, @{const "==>"} $ t1 $ t2) => - do_formula (not pos) t1 andalso - (t2 = @{prop False} orelse do_formula pos t2) - | (_, @{const HOL.implies} $ t1 $ t2) => - do_formula (not pos) t1 andalso - (t2 = @{const False} orelse do_formula pos t2) - | (_, @{const Not} $ t1) => do_formula (not pos) t1 - | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2] - | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2] - | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2 - | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2 - | _ => false - in do_formula true end - -fun has_bound_or_var_of_type tycons = - exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s - | Abs (_, Type (s, _), _) => member (op =) tycons s - | _ => false) - -(* Facts are forbidden to contain variables of these types. The typical reason - is that they lead to unsoundness. Note that "unit" satisfies numerous - equations like "?x = ()". The resulting clauses will have no type constraint, - yielding false proofs. Even "bool" leads to many unsound proofs, though only - for higher-order problems. *) -val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}]; - -(* Facts containing variables of type "unit" or "bool" or of the form - "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types - are omitted. *) -fun is_dangerous_term full_types t = - not full_types andalso - let val t = transform_elim_term t in - has_bound_or_var_of_type dangerous_types t orelse - is_exhaustive_finite t - end - -fun is_theorem_bad_for_atps full_types thm = - let val t = prop_of thm in - is_formula_too_complex t orelse exists_type type_has_top_sort t orelse - is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse - exists_low_level_class_const t orelse is_metastrange_theorem thm orelse - is_that_fact thm - end - -fun clasimpset_rules_of ctxt = - let - val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs - val intros = safeIs @ hazIs - val elims = map Classical.classical_rule (safeEs @ hazEs) - val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd - in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end - -fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths = - let - val thy = ProofContext.theory_of ctxt - val global_facts = PureThy.facts_of thy - val local_facts = ProofContext.facts_of ctxt - val named_locals = local_facts |> Facts.dest_static [] - val is_chained = member Thm.eq_thm chained_ths - val (intros, elims, simps) = - if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then - clasimpset_rules_of ctxt - else - (Termtab.empty, Termtab.empty, Termtab.empty) - (* Unnamed nonchained formulas with schematic variables are omitted, because - they are rejected by the backticks (`...`) parser for some reason. *) - fun is_good_unnamed_local th = - not (Thm.has_name_hint th) andalso - (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso - forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals - val unnamed_locals = - union Thm.eq_thm (Facts.props local_facts) chained_ths - |> filter is_good_unnamed_local |> map (pair "" o single) - val full_space = - Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts) - fun add_facts global foldx facts = - foldx (fn (name0, ths) => - if name0 <> "" andalso - forall (not o member Thm.eq_thm add_thms) ths andalso - (Facts.is_concealed facts name0 orelse - (respect_no_atp andalso is_package_def name0) orelse - exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse - String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then - I - else - let - val multi = length ths > 1 - fun backquotify th = - "`" ^ Print_Mode.setmp [Print_Mode.input] - (Syntax.string_of_term ctxt) (prop_of th) ^ "`" - |> String.translate (fn c => if Char.isPrint c then str c else "") - |> simplify_spaces - fun check_thms a = - case try (ProofContext.get_thms ctxt) a of - NONE => false - | SOME ths' => Thm.eq_thms (ths, ths') - in - pair 1 - #> fold (fn th => fn (j, rest) => - (j + 1, - if is_theorem_bad_for_atps full_types th andalso - not (member Thm.eq_thm add_thms th) then - rest - else - (((fn () => - if name0 = "" then - th |> backquotify - else - let - val name1 = Facts.extern facts name0 - val name2 = Name_Space.extern full_space name0 - in - case find_first check_thms [name1, name2, name0] of - SOME name => repair_name reserved multi j name - | NONE => "" - end), - let val t = prop_of th in - if is_chained th then Chained - else if not global then Local - else if Termtab.defined intros t then Intro - else if Termtab.defined elims t then Elim - else if Termtab.defined simps t then Simp - else General - end), - (multi, th)) :: rest)) ths - #> snd - end) - in - [] |> add_facts false fold local_facts (unnamed_locals @ named_locals) - |> add_facts true Facts.fold_static global_facts global_facts - end - -(* The single-name theorems go after the multiple-name ones, so that single - names are preferred when both are available. *) -fun name_thm_pairs ctxt respect_no_atp = - List.partition (fst o snd) #> op @ #> map (apsnd snd) - #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) - -(***************************************************************) -(* ATP invocation methods setup *) -(***************************************************************) - -fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant - theory_relevant (relevance_override as {add, del, only}) - chained_ths hyp_ts concl_t = - let - val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0), - 1.0 / Real.fromInt (max_relevant + 1)) - val add_thms = maps (ProofContext.get_fact ctxt) add - val reserved = reserved_isar_keyword_table () - val axioms = - (if only then - maps (map (fn ((name, loc), th) => ((K name, loc), (true, th))) - o name_thm_pairs_from_ref ctxt reserved chained_ths) add - else - all_name_thms_pairs ctxt reserved full_types add_thms chained_ths) - |> name_thm_pairs ctxt (respect_no_atp andalso not only) - |> uniquify - in - trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^ - " theorems"); - (if threshold0 > 1.0 orelse threshold0 > threshold1 then - [] - else if threshold0 < 0.0 then - axioms - else - relevance_filter ctxt threshold0 decay max_relevant theory_relevant - relevance_override axioms (concl_t :: hyp_ts)) - |> map (apfst (apfst (fn f => f ()))) - end - -end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML Thu Sep 02 08:40:25 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,162 +0,0 @@ -(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML - Author: Philipp Meyer, TU Muenchen - Author: Jasmin Blanchette, TU Muenchen - -Minimization of theorem list for Metis using automatic theorem provers. -*) - -signature SLEDGEHAMMER_FACT_MINIMIZE = -sig - type locality = Sledgehammer_Fact_Filter.locality - type params = Sledgehammer.params - - val minimize_theorems : - params -> int -> int -> Proof.state -> ((string * locality) * thm list) list - -> ((string * locality) * thm list) list option * string - val run_minimize : params -> int -> Facts.ref list -> Proof.state -> unit -end; - -structure Sledgehammer_Fact_Minimize : SLEDGEHAMMER_FACT_MINIMIZE = -struct - -open ATP_Systems -open Sledgehammer_Util -open Sledgehammer_Fact_Filter -open Sledgehammer_Proof_Reconstruct -open Sledgehammer - -(* wrapper for calling external prover *) - -fun string_for_failure Unprovable = "Unprovable." - | string_for_failure TimedOut = "Timed out." - | string_for_failure _ = "Unknown error." - -fun n_theorems names = - let val n = length names in - string_of_int n ^ " theorem" ^ plural_s n ^ - (if n > 0 then - ": " ^ (names |> map fst - |> sort_distinct string_ord |> space_implode " ") - else - "") - end - -fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof, - isar_shrink_factor, ...} : params) - (prover : prover) explicit_apply timeout subgoal state - axioms = - let - val _ = - priority ("Testing " ^ n_theorems (map fst axioms) ^ "...") - val params = - {debug = debug, verbose = verbose, overlord = overlord, atps = atps, - full_types = full_types, explicit_apply = explicit_apply, - relevance_thresholds = (1.01, 1.01), max_relevant = NONE, - theory_relevant = NONE, isar_proof = isar_proof, - isar_shrink_factor = isar_shrink_factor, timeout = timeout} - val axioms = maps (fn (n, ths) => map (pair n) ths) axioms - val {context = ctxt, facts, goal} = Proof.goal state - val problem = - {subgoal = subgoal, goal = (ctxt, (facts, goal)), - relevance_override = {add = [], del = [], only = false}, - axioms = SOME axioms} - val result as {outcome, used_thm_names, ...} = prover params (K "") problem - in - priority (case outcome of - NONE => - if length used_thm_names = length axioms then - "Found proof." - else - "Found proof with " ^ n_theorems used_thm_names ^ "." - | SOME failure => string_for_failure failure); - result - end - -(* minimalization of thms *) - -fun filter_used_facts used = filter (member (op =) used o fst) - -fun sublinear_minimize _ [] p = p - | sublinear_minimize test (x :: xs) (seen, result) = - case test (xs @ seen) of - result as {outcome = NONE, proof, used_thm_names, ...} : prover_result => - sublinear_minimize test (filter_used_facts used_thm_names xs) - (filter_used_facts used_thm_names seen, result) - | _ => sublinear_minimize test xs (x :: seen, result) - -(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer - preprocessing time is included in the estimate below but isn't part of the - timeout. *) -val fudge_msecs = 1000 - -fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set." - | minimize_theorems (params as {debug, atps = atp :: _, full_types, - isar_proof, isar_shrink_factor, timeout, ...}) - i n state axioms = - let - val thy = Proof.theory_of state - val prover = get_prover_fun thy atp - val msecs = Time.toMilliseconds timeout - val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".") - val {context = ctxt, goal, ...} = Proof.goal state - val (_, hyp_ts, concl_t) = strip_subgoal goal i - val explicit_apply = - not (forall (Meson.is_fol_term thy) - (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms)) - fun do_test timeout = - test_theorems params prover explicit_apply timeout i state - val timer = Timer.startRealTimer () - in - (case do_test timeout axioms of - result as {outcome = NONE, pool, used_thm_names, - conjecture_shape, ...} => - let - val time = Timer.checkRealTimer timer - val new_timeout = - Int.min (msecs, Time.toMilliseconds time + fudge_msecs) - |> Time.fromMilliseconds - val (min_thms, {proof, axiom_names, ...}) = - sublinear_minimize (do_test new_timeout) - (filter_used_facts used_thm_names axioms) ([], result) - val n = length min_thms - val _ = priority (cat_lines - ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^ - (case length (filter (curry (op =) Chained o snd o fst) min_thms) of - 0 => "" - | n => " (including " ^ Int.toString n ^ " chained)") ^ ".") - in - (SOME min_thms, - proof_text isar_proof - (pool, debug, isar_shrink_factor, ctxt, conjecture_shape) - (full_types, K "", proof, axiom_names, goal, i) |> fst) - end - | {outcome = SOME TimedOut, ...} => - (NONE, "Timeout: You can increase the time limit using the \"timeout\" \ - \option (e.g., \"timeout = " ^ - string_of_int (10 + msecs div 1000) ^ " s\").") - | {outcome = SOME UnknownError, ...} => - (* Failure sometimes mean timeout, unfortunately. *) - (NONE, "Failure: No proof was found with the current time limit. You \ - \can increase the time limit using the \"timeout\" \ - \option (e.g., \"timeout = " ^ - string_of_int (10 + msecs div 1000) ^ " s\").") - | {message, ...} => (NONE, "ATP error: " ^ message)) - handle ERROR msg => (NONE, "Error: " ^ msg) - end - -fun run_minimize params i refs state = - let - val ctxt = Proof.context_of state - val reserved = reserved_isar_keyword_table () - val chained_ths = #facts (Proof.goal state) - val axioms = - maps (map (apsnd single) - o name_thm_pairs_from_ref ctxt reserved chained_ths) refs - in - case subgoal_count state of - 0 => priority "No subgoal!" - | n => - (kill_atps (); priority (#2 (minimize_theorems params i n state axioms))) - end - -end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML Thu Sep 02 08:51:16 2010 +0200 @@ -0,0 +1,812 @@ +(* Title: HOL/Tools/Sledgehammer/sledgehammer_filter.ML + Author: Jia Meng, Cambridge University Computer Laboratory and NICTA + Author: Jasmin Blanchette, TU Muenchen +*) + +signature SLEDGEHAMMER_FILTER = +sig + datatype locality = General | Intro | Elim | Simp | Local | Assum | Chained + + type relevance_override = + {add: (Facts.ref * Attrib.src list) list, + del: (Facts.ref * Attrib.src list) list, + only: bool} + + val trace : bool Unsynchronized.ref + val worse_irrel_freq : real Unsynchronized.ref + val higher_order_irrel_weight : real Unsynchronized.ref + val abs_rel_weight : real Unsynchronized.ref + val abs_irrel_weight : real Unsynchronized.ref + val skolem_irrel_weight : real Unsynchronized.ref + val theory_const_rel_weight : real Unsynchronized.ref + val theory_const_irrel_weight : real Unsynchronized.ref + val intro_bonus : real Unsynchronized.ref + val elim_bonus : real Unsynchronized.ref + val simp_bonus : real Unsynchronized.ref + val local_bonus : real Unsynchronized.ref + val assum_bonus : real Unsynchronized.ref + val chained_bonus : real Unsynchronized.ref + val max_imperfect : real Unsynchronized.ref + val max_imperfect_exp : real Unsynchronized.ref + val threshold_divisor : real Unsynchronized.ref + val ridiculous_threshold : real Unsynchronized.ref + val name_thm_pairs_from_ref : + Proof.context -> unit Symtab.table -> thm list + -> Facts.ref * Attrib.src list -> ((string * locality) * thm) list + val relevant_facts : + Proof.context -> bool -> real * real -> int -> relevance_override + -> thm list -> term list -> term -> ((string * locality) * thm) list +end; + +structure Sledgehammer_Filter : SLEDGEHAMMER_FILTER = +struct + +open Sledgehammer_Util + +val trace = Unsynchronized.ref false +fun trace_msg msg = if !trace then tracing (msg ()) else () + +(* experimental features *) +val term_patterns = false +val respect_no_atp = true + +(* FUDGE *) +val worse_irrel_freq = Unsynchronized.ref 100.0 +val higher_order_irrel_weight = Unsynchronized.ref 1.05 +val abs_rel_weight = Unsynchronized.ref 0.5 +val abs_irrel_weight = Unsynchronized.ref 2.0 +val skolem_irrel_weight = Unsynchronized.ref 0.75 +val theory_const_rel_weight = Unsynchronized.ref 0.5 +val theory_const_irrel_weight = Unsynchronized.ref 0.25 +val intro_bonus = Unsynchronized.ref 0.15 +val elim_bonus = Unsynchronized.ref 0.15 +val simp_bonus = Unsynchronized.ref 0.15 +val local_bonus = Unsynchronized.ref 0.55 +val assum_bonus = Unsynchronized.ref 1.05 +val chained_bonus = Unsynchronized.ref 1.5 +val max_imperfect = Unsynchronized.ref 11.5 +val max_imperfect_exp = Unsynchronized.ref 1.0 +val threshold_divisor = Unsynchronized.ref 2.0 +val ridiculous_threshold = Unsynchronized.ref 0.1 + +datatype locality = General | Intro | Elim | Simp | Local | Assum | Chained + +type relevance_override = + {add: (Facts.ref * Attrib.src list) list, + del: (Facts.ref * Attrib.src list) list, + only: bool} + +val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator +val abs_name = "Sledgehammer.abs" +val skolem_prefix = "Sledgehammer.sko" +val theory_const_suffix = Long_Name.separator ^ " 1" + +fun repair_name reserved multi j name = + (name |> Symtab.defined reserved name ? quote) ^ + (if multi then "(" ^ Int.toString j ^ ")" else "") + +fun name_thm_pairs_from_ref ctxt reserved chained_ths (xthm as (xref, args)) = + let + val ths = Attrib.eval_thms ctxt [xthm] + val bracket = + implode (map (fn arg => "[" ^ Pretty.str_of (Args.pretty_src ctxt arg) + ^ "]") args) + val space_bracket = if bracket = "" then "" else " " ^ bracket + val name = + case xref of + Facts.Fact s => "`" ^ s ^ "`" ^ space_bracket + | Facts.Named (("", _), _) => "[" ^ bracket ^ "]" + | _ => Facts.string_of_ref xref ^ space_bracket + val multi = length ths > 1 + in + (ths, (1, [])) + |-> fold (fn th => fn (j, rest) => + (j + 1, ((repair_name reserved multi j name, + if member Thm.eq_thm chained_ths th then Chained + else General), th) :: rest)) + |> snd + end + +(***************************************************************) +(* Relevance Filtering *) +(***************************************************************) + +(*** constants with types ***) + +fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) = + order_of_type T1 (* cheat: pretend sets are first-order *) + | order_of_type (Type (@{type_name fun}, [T1, T2])) = + Int.max (order_of_type T1 + 1, order_of_type T2) + | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0 + | order_of_type _ = 0 + +(* An abstraction of Isabelle types and first-order terms *) +datatype pattern = PVar | PApp of string * pattern list +datatype ptype = PType of int * pattern list + +fun string_for_pattern PVar = "_" + | string_for_pattern (PApp (s, ps)) = + if null ps then s else s ^ string_for_patterns ps +and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")" +fun string_for_ptype (PType (_, ps)) = string_for_patterns ps + +(*Is the second type an instance of the first one?*) +fun match_pattern (PVar, _) = true + | match_pattern (PApp _, PVar) = false + | match_pattern (PApp (s, ps), PApp (t, qs)) = + s = t andalso match_patterns (ps, qs) +and match_patterns (_, []) = true + | match_patterns ([], _) = false + | match_patterns (p :: ps, q :: qs) = + match_pattern (p, q) andalso match_patterns (ps, qs) +fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs) + +(* Is there a unifiable constant? *) +fun pconst_mem f consts (s, ps) = + exists (curry (match_ptype o f) ps) + (map snd (filter (curry (op =) s o fst) consts)) +fun pconst_hyper_mem f const_tab (s, ps) = + exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s)) + +fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts) + | pattern_for_type (TFree (s, _)) = PApp (s, []) + | pattern_for_type (TVar _) = PVar + +fun pterm thy t = + case strip_comb t of + (Const x, ts) => PApp (pconst thy true x ts) + | (Free x, ts) => PApp (pconst thy false x ts) + | (Var x, []) => PVar + | _ => PApp ("?", []) (* equivalence class of higher-order constructs *) +(* Pairs a constant with the list of its type instantiations. *) +and ptype thy const x ts = + (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x)) + else []) @ + (if term_patterns then map (pterm thy) ts else []) +and pconst thy const (s, T) ts = (s, ptype thy const (s, T) ts) +and rich_ptype thy const (s, T) ts = + PType (order_of_type T, ptype thy const (s, T) ts) +and rich_pconst thy const (s, T) ts = (s, rich_ptype thy const (s, T) ts) + +fun string_for_hyper_pconst (s, ps) = + s ^ "{" ^ commas (map string_for_ptype ps) ^ "}" + +(* These are typically simplified away by "Meson.presimplify". Equality is + handled specially via "fequal". *) +val boring_consts = + [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let}, + @{const_name HOL.eq}] + +(* Add a pconstant to the table, but a [] entry means a standard + connective, which we ignore.*) +fun add_pconst_to_table also_skolem (c, p) = + if member (op =) boring_consts c orelse + (not also_skolem andalso String.isPrefix skolem_prefix c) then + I + else + Symtab.map_default (c, [p]) (insert (op =) p) + +fun is_formula_type T = (T = HOLogic.boolT orelse T = propT) + +fun pconsts_in_terms thy also_skolems pos ts = + let + val flip = Option.map not + (* We include free variables, as well as constants, to handle locales. For + each quantifiers that must necessarily be skolemized by the ATP, we + introduce a fresh constant to simulate the effect of Skolemization. *) + fun do_const const (s, T) ts = + add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts) + #> fold do_term ts + and do_term t = + case strip_comb t of + (Const x, ts) => do_const true x ts + | (Free x, ts) => do_const false x ts + | (Abs (_, T, t'), ts) => + (null ts + ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, []))) + #> fold do_term (t' :: ts) + | (_, ts) => fold do_term ts + fun do_quantifier will_surely_be_skolemized abs_T body_t = + do_formula pos body_t + #> (if also_skolems andalso will_surely_be_skolemized then + add_pconst_to_table true + (gensym skolem_prefix, PType (order_of_type abs_T, [])) + else + I) + and do_term_or_formula T = + if is_formula_type T then do_formula NONE else do_term + and do_formula pos t = + case t of + Const (@{const_name all}, _) $ Abs (_, T, t') => + do_quantifier (pos = SOME false) T t' + | @{const "==>"} $ t1 $ t2 => + do_formula (flip pos) t1 #> do_formula pos t2 + | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 => + fold (do_term_or_formula T) [t1, t2] + | @{const Trueprop} $ t1 => do_formula pos t1 + | @{const Not} $ t1 => do_formula (flip pos) t1 + | Const (@{const_name All}, _) $ Abs (_, T, t') => + do_quantifier (pos = SOME false) T t' + | Const (@{const_name Ex}, _) $ Abs (_, T, t') => + do_quantifier (pos = SOME true) T t' + | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2] + | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2] + | @{const HOL.implies} $ t1 $ t2 => + do_formula (flip pos) t1 #> do_formula pos t2 + | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 => + fold (do_term_or_formula T) [t1, t2] + | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])])) + $ t1 $ t2 $ t3 => + do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3] + | Const (@{const_name Ex1}, _) $ Abs (_, T, t') => + do_quantifier (is_some pos) T t' + | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') => + do_quantifier (pos = SOME false) T + (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t')) + | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') => + do_quantifier (pos = SOME true) T + (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t')) + | (t0 as Const (_, @{typ bool})) $ t1 => + do_term t0 #> do_formula pos t1 (* theory constant *) + | _ => do_term t + in Symtab.empty |> fold (do_formula pos) ts end + +(*Inserts a dummy "constant" referring to the theory name, so that relevance + takes the given theory into account.*) +fun theory_const_prop_of th = + if exists (curry (op <) 0.0) [!theory_const_rel_weight, + !theory_const_irrel_weight] then + let + val name = Context.theory_name (theory_of_thm th) + val t = Const (name ^ theory_const_suffix, @{typ bool}) + in t $ prop_of th end + else + prop_of th + +(**** Constant / Type Frequencies ****) + +(* A two-dimensional symbol table counts frequencies of constants. It's keyed + first by constant name and second by its list of type instantiations. For the + latter, we need a linear ordering on "pattern list". *) + +fun pattern_ord p = + case p of + (PVar, PVar) => EQUAL + | (PVar, PApp _) => LESS + | (PApp _, PVar) => GREATER + | (PApp q1, PApp q2) => + prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2) +fun ptype_ord (PType p, PType q) = + prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q) + +structure PType_Tab = Table(type key = ptype val ord = ptype_ord) + +fun count_axiom_consts thy = + let + fun do_const const (s, T) ts = + (* Two-dimensional table update. Constant maps to types maps to count. *) + PType_Tab.map_default (rich_ptype thy const (s, T) ts, 0) (Integer.add 1) + |> Symtab.map_default (s, PType_Tab.empty) + #> fold do_term ts + and do_term t = + case strip_comb t of + (Const x, ts) => do_const true x ts + | (Free x, ts) => do_const false x ts + | (Abs (_, _, t'), ts) => fold do_term (t' :: ts) + | (_, ts) => fold do_term ts + in do_term o theory_const_prop_of o snd end + + +(**** Actual Filtering Code ****) + +fun pow_int x 0 = 1.0 + | pow_int x 1 = x + | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x + +(*The frequency of a constant is the sum of those of all instances of its type.*) +fun pconst_freq match const_tab (c, ps) = + PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m) + (the (Symtab.lookup const_tab c)) 0 + + +(* A surprising number of theorems contain only a few significant constants. + These include all induction rules, and other general theorems. *) + +(* "log" seems best in practice. A constant function of one ignores the constant + frequencies. Rare constants give more points if they are relevant than less + rare ones. *) +fun rel_weight_for order freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0) + +(* Irrelevant constants are treated differently. We associate lower penalties to + very rare constants and very common ones -- the former because they can't + lead to the inclusion of too many new facts, and the latter because they are + so common as to be of little interest. *) +fun irrel_weight_for order freq = + let val (k, x) = !worse_irrel_freq |> `Real.ceil in + (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x + else rel_weight_for order freq / rel_weight_for order k) + * pow_int (!higher_order_irrel_weight) (order - 1) + end + +(* Computes a constant's weight, as determined by its frequency. *) +fun generic_pconst_weight abs_weight skolem_weight theory_const_weight + weight_for f const_tab (c as (s, PType (m, _))) = + if s = abs_name then abs_weight + else if String.isPrefix skolem_prefix s then skolem_weight + else if String.isSuffix theory_const_suffix s then theory_const_weight + else weight_for m (pconst_freq (match_ptype o f) const_tab c) + +fun rel_pconst_weight const_tab = + generic_pconst_weight (!abs_rel_weight) 0.0 (!theory_const_rel_weight) + rel_weight_for I const_tab +fun irrel_pconst_weight const_tab = + generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight) + (!theory_const_irrel_weight) irrel_weight_for swap const_tab + +fun locality_bonus General = 0.0 + | locality_bonus Intro = !intro_bonus + | locality_bonus Elim = !elim_bonus + | locality_bonus Simp = !simp_bonus + | locality_bonus Local = !local_bonus + | locality_bonus Assum = !assum_bonus + | locality_bonus Chained = !chained_bonus + +fun axiom_weight loc const_tab relevant_consts axiom_consts = + case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts) + ||> filter_out (pconst_hyper_mem swap relevant_consts) of + ([], _) => 0.0 + | (rel, irrel) => + let + val irrel = irrel |> filter_out (pconst_mem swap rel) + val rel_weight = + 0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel + val irrel_weight = + ~ (locality_bonus loc) + |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel + val res = rel_weight / (rel_weight + irrel_weight) + in if Real.isFinite res then res else 0.0 end + +(* FIXME: experiment +fun debug_axiom_weight loc const_tab relevant_consts axiom_consts = + case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts) + ||> filter_out (pconst_hyper_mem swap relevant_consts) of + ([], _) => 0.0 + | (rel, irrel) => + let + val irrel = irrel |> filter_out (pconst_mem swap rel) + val rels_weight = + 0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel + val irrels_weight = + ~ (locality_bonus loc) + |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel +val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel)) +val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel)) + val res = rels_weight / (rels_weight + irrels_weight) + in if Real.isFinite res then res else 0.0 end +*) + +fun pconsts_in_axiom thy t = + Symtab.fold (fn (s, pss) => fold (cons o pair s) pss) + (pconsts_in_terms thy true (SOME true) [t]) [] +fun pair_consts_axiom thy axiom = + case axiom |> snd |> theory_const_prop_of |> pconsts_in_axiom thy of + [] => NONE + | consts => SOME ((axiom, consts), NONE) + +type annotated_thm = + (((unit -> string) * locality) * thm) * (string * ptype) list + +fun take_most_relevant max_relevant remaining_max + (candidates : (annotated_thm * real) list) = + let + val max_imperfect = + Real.ceil (Math.pow (!max_imperfect, + Math.pow (Real.fromInt remaining_max + / Real.fromInt max_relevant, !max_imperfect_exp))) + val (perfect, imperfect) = + candidates |> sort (Real.compare o swap o pairself snd) + |> take_prefix (fn (_, w) => w > 0.99999) + val ((accepts, more_rejects), rejects) = + chop max_imperfect imperfect |>> append perfect |>> chop remaining_max + in + trace_msg (fn () => + "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^ + Int.toString (length candidates) ^ "): " ^ + (accepts |> map (fn ((((name, _), _), _), weight) => + name () ^ " [" ^ Real.toString weight ^ "]") + |> commas)); + (accepts, more_rejects @ rejects) + end + +fun if_empty_replace_with_locality thy axioms loc tab = + if Symtab.is_empty tab then + pconsts_in_terms thy false (SOME false) + (map_filter (fn ((_, loc'), th) => + if loc' = loc then SOME (prop_of th) else NONE) axioms) + else + tab + +fun relevance_filter ctxt threshold0 decay max_relevant + ({add, del, ...} : relevance_override) axioms goal_ts = + let + val thy = ProofContext.theory_of ctxt + val const_tab = fold (count_axiom_consts thy) axioms Symtab.empty + val goal_const_tab = + pconsts_in_terms thy false (SOME false) goal_ts + |> fold (if_empty_replace_with_locality thy axioms) + [Chained, Assum, Local] + val add_ths = Attrib.eval_thms ctxt add + val del_ths = Attrib.eval_thms ctxt del + fun iter j remaining_max threshold rel_const_tab hopeless hopeful = + let + fun game_over rejects = + (* Add "add:" facts. *) + if null add_ths then + [] + else + map_filter (fn ((p as (_, th), _), _) => + if member Thm.eq_thm add_ths th then SOME p + else NONE) rejects + fun relevant [] rejects [] = + (* Nothing has been added this iteration. *) + if j = 0 andalso threshold >= !ridiculous_threshold then + (* First iteration? Try again. *) + iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab + hopeless hopeful + else + game_over (rejects @ hopeless) + | relevant candidates rejects [] = + let + val (accepts, more_rejects) = + take_most_relevant max_relevant remaining_max candidates + val rel_const_tab' = + rel_const_tab + |> fold (add_pconst_to_table false) (maps (snd o fst) accepts) + fun is_dirty (c, _) = + Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c + val (hopeful_rejects, hopeless_rejects) = + (rejects @ hopeless, ([], [])) + |-> fold (fn (ax as (_, consts), old_weight) => + if exists is_dirty consts then + apfst (cons (ax, NONE)) + else + apsnd (cons (ax, old_weight))) + |>> append (more_rejects + |> map (fn (ax as (_, consts), old_weight) => + (ax, if exists is_dirty consts then NONE + else SOME old_weight))) + val threshold = + 1.0 - (1.0 - threshold) + * Math.pow (decay, Real.fromInt (length accepts)) + val remaining_max = remaining_max - length accepts + in + trace_msg (fn () => "New or updated constants: " ^ + commas (rel_const_tab' |> Symtab.dest + |> subtract (op =) (rel_const_tab |> Symtab.dest) + |> map string_for_hyper_pconst)); + map (fst o fst) accepts @ + (if remaining_max = 0 then + game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects) + else + iter (j + 1) remaining_max threshold rel_const_tab' + hopeless_rejects hopeful_rejects) + end + | relevant candidates rejects + (((ax as (((_, loc), th), axiom_consts)), cached_weight) + :: hopeful) = + let + val weight = + case cached_weight of + SOME w => w + | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts +(* FIXME: experiment +val name = fst (fst (fst ax)) () +val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then +tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts)) +else +() +*) + in + if weight >= threshold then + relevant ((ax, weight) :: candidates) rejects hopeful + else + relevant candidates ((ax, weight) :: rejects) hopeful + end + in + trace_msg (fn () => + "ITERATION " ^ string_of_int j ^ ": current threshold: " ^ + Real.toString threshold ^ ", constants: " ^ + commas (rel_const_tab |> Symtab.dest + |> filter (curry (op <>) [] o snd) + |> map string_for_hyper_pconst)); + relevant [] [] hopeful + end + in + axioms |> filter_out (member Thm.eq_thm del_ths o snd) + |> map_filter (pair_consts_axiom thy) + |> iter 0 max_relevant threshold0 goal_const_tab [] + |> tap (fn res => trace_msg (fn () => + "Total relevant: " ^ Int.toString (length res))) + end + + +(***************************************************************) +(* Retrieving and filtering lemmas *) +(***************************************************************) + +(*** retrieve lemmas and filter them ***) + +(*Reject theorems with names like "List.filter.filter_list_def" or + "Accessible_Part.acc.defs", as these are definitions arising from packages.*) +fun is_package_def a = + let val names = Long_Name.explode a + in + length names > 2 andalso + not (hd names = "local") andalso + String.isSuffix "_def" a orelse String.isSuffix "_defs" a + end; + +fun mk_fact_table f xs = + fold (Termtab.update o `(prop_of o f)) xs Termtab.empty +fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) [] + +(* FIXME: put other record thms here, or declare as "no_atp" *) +val multi_base_blacklist = + ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits", + "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy", + "case_cong", "weak_case_cong"] + |> map (prefix ".") + +val max_lambda_nesting = 3 + +fun term_has_too_many_lambdas max (t1 $ t2) = + exists (term_has_too_many_lambdas max) [t1, t2] + | term_has_too_many_lambdas max (Abs (_, _, t)) = + max = 0 orelse term_has_too_many_lambdas (max - 1) t + | term_has_too_many_lambdas _ _ = false + +(* Don't count nested lambdas at the level of formulas, since they are + quantifiers. *) +fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) = + formula_has_too_many_lambdas (T :: Ts) t + | formula_has_too_many_lambdas Ts t = + if is_formula_type (fastype_of1 (Ts, t)) then + exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t)) + else + term_has_too_many_lambdas max_lambda_nesting t + +(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31) + was 11. *) +val max_apply_depth = 15 + +fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1) + | apply_depth (Abs (_, _, t)) = apply_depth t + | apply_depth _ = 0 + +fun is_formula_too_complex t = + apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t + +val exists_sledgehammer_const = + exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s) + +(* FIXME: make more reliable *) +val exists_low_level_class_const = + exists_Const (fn (s, _) => + String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s) + +fun is_metastrange_theorem th = + case head_of (concl_of th) of + Const (a, _) => (a <> @{const_name Trueprop} andalso + a <> @{const_name "=="}) + | _ => false + +fun is_that_fact th = + String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th) + andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN + | _ => false) (prop_of th) + +val type_has_top_sort = + exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false) + +(**** Predicates to detect unwanted facts (prolific or likely to cause + unsoundness) ****) + +(* Too general means, positive equality literal with a variable X as one + operand, when X does not occur properly in the other operand. This rules out + clearly inconsistent facts such as X = a | X = b, though it by no means + guarantees soundness. *) + +(* Unwanted equalities are those between a (bound or schematic) variable that + does not properly occur in the second operand. *) +val is_exhaustive_finite = + let + fun is_bad_equal (Var z) t = + not (exists_subterm (fn Var z' => z = z' | _ => false) t) + | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j)) + | is_bad_equal _ _ = false + fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1 + fun do_formula pos t = + case (pos, t) of + (_, @{const Trueprop} $ t1) => do_formula pos t1 + | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) => + do_formula pos t' + | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) => + do_formula pos t' + | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) => + do_formula pos t' + | (_, @{const "==>"} $ t1 $ t2) => + do_formula (not pos) t1 andalso + (t2 = @{prop False} orelse do_formula pos t2) + | (_, @{const HOL.implies} $ t1 $ t2) => + do_formula (not pos) t1 andalso + (t2 = @{const False} orelse do_formula pos t2) + | (_, @{const Not} $ t1) => do_formula (not pos) t1 + | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2] + | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2] + | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2 + | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2 + | _ => false + in do_formula true end + +fun has_bound_or_var_of_type tycons = + exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s + | Abs (_, Type (s, _), _) => member (op =) tycons s + | _ => false) + +(* Facts are forbidden to contain variables of these types. The typical reason + is that they lead to unsoundness. Note that "unit" satisfies numerous + equations like "?x = ()". The resulting clauses will have no type constraint, + yielding false proofs. Even "bool" leads to many unsound proofs, though only + for higher-order problems. *) +val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}]; + +(* Facts containing variables of type "unit" or "bool" or of the form + "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types + are omitted. *) +fun is_dangerous_term full_types t = + not full_types andalso + let val t = transform_elim_term t in + has_bound_or_var_of_type dangerous_types t orelse + is_exhaustive_finite t + end + +fun is_theorem_bad_for_atps full_types thm = + let val t = prop_of thm in + is_formula_too_complex t orelse exists_type type_has_top_sort t orelse + is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse + exists_low_level_class_const t orelse is_metastrange_theorem thm orelse + is_that_fact thm + end + +fun clasimpset_rules_of ctxt = + let + val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs + val intros = safeIs @ hazIs + val elims = map Classical.classical_rule (safeEs @ hazEs) + val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd + in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end + +fun all_name_thms_pairs ctxt reserved full_types add_ths chained_ths = + let + val thy = ProofContext.theory_of ctxt + val global_facts = PureThy.facts_of thy + val local_facts = ProofContext.facts_of ctxt + val named_locals = local_facts |> Facts.dest_static [] + val assms = Assumption.all_assms_of ctxt + fun is_assum th = exists (fn ct => prop_of th aconv term_of ct) assms + val is_chained = member Thm.eq_thm chained_ths + val (intros, elims, simps) = + if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then + clasimpset_rules_of ctxt + else + (Termtab.empty, Termtab.empty, Termtab.empty) + (* Unnamed nonchained formulas with schematic variables are omitted, because + they are rejected by the backticks (`...`) parser for some reason. *) + fun is_good_unnamed_local th = + not (Thm.has_name_hint th) andalso + (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso + forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals + val unnamed_locals = + union Thm.eq_thm (Facts.props local_facts) chained_ths + |> filter is_good_unnamed_local |> map (pair "" o single) + val full_space = + Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts) + fun add_facts global foldx facts = + foldx (fn (name0, ths) => + if name0 <> "" andalso + forall (not o member Thm.eq_thm add_ths) ths andalso + (Facts.is_concealed facts name0 orelse + (respect_no_atp andalso is_package_def name0) orelse + exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse + String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then + I + else + let + val multi = length ths > 1 + fun backquotify th = + "`" ^ Print_Mode.setmp [Print_Mode.input] + (Syntax.string_of_term ctxt) (prop_of th) ^ "`" + |> String.translate (fn c => if Char.isPrint c then str c else "") + |> simplify_spaces + fun check_thms a = + case try (ProofContext.get_thms ctxt) a of + NONE => false + | SOME ths' => Thm.eq_thms (ths, ths') + in + pair 1 + #> fold (fn th => fn (j, rest) => + (j + 1, + if is_theorem_bad_for_atps full_types th andalso + not (member Thm.eq_thm add_ths th) then + rest + else + (((fn () => + if name0 = "" then + th |> backquotify + else + let + val name1 = Facts.extern facts name0 + val name2 = Name_Space.extern full_space name0 + in + case find_first check_thms [name1, name2, name0] of + SOME name => repair_name reserved multi j name + | NONE => "" + end), + let val t = prop_of th in + if is_chained th then Chained + else if global then + if Termtab.defined intros t then Intro + else if Termtab.defined elims t then Elim + else if Termtab.defined simps t then Simp + else General + else + if is_assum th then Assum else Local + end), + (multi, th)) :: rest)) ths + #> snd + end) + in + [] |> add_facts false fold local_facts (unnamed_locals @ named_locals) + |> add_facts true Facts.fold_static global_facts global_facts + end + +(* The single-name theorems go after the multiple-name ones, so that single + names are preferred when both are available. *) +fun name_thm_pairs ctxt respect_no_atp = + List.partition (fst o snd) #> op @ #> map (apsnd snd) + #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) + +(***************************************************************) +(* ATP invocation methods setup *) +(***************************************************************) + +fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant + (relevance_override as {add, del, only}) chained_ths hyp_ts + concl_t = + let + val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0), + 1.0 / Real.fromInt (max_relevant + 1)) + val add_ths = Attrib.eval_thms ctxt add + val reserved = reserved_isar_keyword_table () + val axioms = + (if only then + maps (map (fn ((name, loc), th) => ((K name, loc), (true, th))) + o name_thm_pairs_from_ref ctxt reserved chained_ths) add + else + all_name_thms_pairs ctxt reserved full_types add_ths chained_ths) + |> name_thm_pairs ctxt (respect_no_atp andalso not only) + |> uniquify + in + trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^ + " theorems"); + (if threshold0 > 1.0 orelse threshold0 > threshold1 then + [] + else if threshold0 < 0.0 then + axioms + else + relevance_filter ctxt threshold0 decay max_relevant relevance_override + axioms (concl_t :: hyp_ts)) + |> map (apfst (apfst (fn f => f ()))) + end + +end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Thu Sep 02 08:51:16 2010 +0200 @@ -20,7 +20,7 @@ open ATP_Systems open Sledgehammer_Util open Sledgehammer -open Sledgehammer_Fact_Minimize +open Sledgehammer_Minimize (** Sledgehammer commands **) @@ -63,25 +63,25 @@ type raw_param = string * string list val default_default_params = - [("debug", "false"), + [("blocking", "false"), + ("debug", "false"), ("verbose", "false"), ("overlord", "false"), ("explicit_apply", "false"), ("relevance_thresholds", "45 85"), ("max_relevant", "smart"), - ("theory_relevant", "smart"), ("isar_proof", "false"), ("isar_shrink_factor", "1")] val alias_params = [("atp", "atps")] val negated_alias_params = - [("no_debug", "debug"), + [("non_blocking", "blocking"), + ("no_debug", "debug"), ("quiet", "verbose"), ("no_overlord", "overlord"), ("partial_types", "full_types"), ("implicit_apply", "explicit_apply"), - ("theory_irrelevant", "theory_relevant"), ("no_isar_proof", "isar_proof")] val params_for_minimize = @@ -94,7 +94,7 @@ AList.defined (op =) default_default_params s orelse AList.defined (op =) alias_params s orelse AList.defined (op =) negated_alias_params s orelse - member (op =) property_dependent_params s + member (op =) property_dependent_params s orelse s = "expect" fun check_raw_param (s, _) = if is_known_raw_param s then () @@ -167,6 +167,7 @@ case lookup name of SOME "smart" => NONE | _ => SOME (lookup_int name) + val blocking = lookup_bool "blocking" val debug = lookup_bool "debug" val verbose = debug orelse lookup_bool "verbose" val overlord = lookup_bool "overlord" @@ -177,16 +178,16 @@ lookup_int_pair "relevance_thresholds" |> pairself (fn n => 0.01 * Real.fromInt n) val max_relevant = lookup_int_option "max_relevant" - val theory_relevant = lookup_bool_option "theory_relevant" val isar_proof = lookup_bool "isar_proof" val isar_shrink_factor = Int.max (1, lookup_int "isar_shrink_factor") val timeout = lookup_time "timeout" + val expect = lookup_string "expect" in - {debug = debug, verbose = verbose, overlord = overlord, atps = atps, - full_types = full_types, explicit_apply = explicit_apply, + {blocking = blocking, debug = debug, verbose = verbose, overlord = overlord, + atps = atps, full_types = full_types, explicit_apply = explicit_apply, relevance_thresholds = relevance_thresholds, max_relevant = max_relevant, - theory_relevant = theory_relevant, isar_proof = isar_proof, - isar_shrink_factor = isar_shrink_factor, timeout = timeout} + isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor, + timeout = timeout, expect = expect} end fun get_params thy = extract_params (default_raw_params thy) @@ -267,10 +268,7 @@ val parse_param = parse_key -- Scan.optional (Parse.$$$ "=" |-- parse_value) [] val parse_params = Scan.optional (Args.bracks (Parse.list parse_param)) [] val parse_fact_refs = - Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) - (Parse.xname -- Scan.option Attrib.thm_sel) - >> (fn (name, interval) => - Facts.Named ((name, Position.none), interval))) + Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse_Spec.xthm) val parse_relevance_chunk = (Args.add |-- Args.colon |-- parse_fact_refs >> add_to_relevance_override) || (Args.del |-- Args.colon |-- parse_fact_refs diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Thu Sep 02 08:51:16 2010 +0200 @@ -0,0 +1,163 @@ +(* Title: HOL/Tools/Sledgehammer/sledgehammer_minimize.ML + Author: Philipp Meyer, TU Muenchen + Author: Jasmin Blanchette, TU Muenchen + +Minimization of theorem list for Metis using automatic theorem provers. +*) + +signature SLEDGEHAMMER_MINIMIZE = +sig + type locality = Sledgehammer_Filter.locality + type params = Sledgehammer.params + + val minimize_theorems : + params -> int -> int -> Proof.state -> ((string * locality) * thm list) list + -> ((string * locality) * thm list) list option * string + val run_minimize : + params -> int -> (Facts.ref * Attrib.src list) list -> Proof.state -> unit +end; + +structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE = +struct + +open ATP_Systems +open Sledgehammer_Util +open Sledgehammer_Filter +open Sledgehammer_Translate +open Sledgehammer_Reconstruct +open Sledgehammer + +(* wrapper for calling external prover *) + +fun string_for_failure Unprovable = "Unprovable." + | string_for_failure TimedOut = "Timed out." + | string_for_failure _ = "Unknown error." + +fun n_theorems names = + let val n = length names in + string_of_int n ^ " theorem" ^ plural_s n ^ + (if n > 0 then + ": " ^ (names |> map fst + |> sort_distinct string_ord |> space_implode " ") + else + "") + end + +fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof, + isar_shrink_factor, ...} : params) + (prover : prover) explicit_apply timeout subgoal state + axioms = + let + val _ = + priority ("Testing " ^ n_theorems (map fst axioms) ^ "...") + val params = + {blocking = true, debug = debug, verbose = verbose, overlord = overlord, + atps = atps, full_types = full_types, explicit_apply = explicit_apply, + relevance_thresholds = (1.01, 1.01), + max_relevant = SOME 65536 (* a large number *), isar_proof = isar_proof, + isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""} + val axioms = maps (fn (n, ths) => map (pair n) ths) axioms + val {context = ctxt, goal, ...} = Proof.goal state + val problem = + {ctxt = ctxt, goal = goal, subgoal = subgoal, + axioms = map (prepare_axiom ctxt) axioms} + val result as {outcome, used_thm_names, ...} = prover params (K "") problem + in + priority (case outcome of + NONE => + if length used_thm_names = length axioms then + "Found proof." + else + "Found proof with " ^ n_theorems used_thm_names ^ "." + | SOME failure => string_for_failure failure); + result + end + +(* minimalization of thms *) + +fun filter_used_facts used = filter (member (op =) used o fst) + +fun sublinear_minimize _ [] p = p + | sublinear_minimize test (x :: xs) (seen, result) = + case test (xs @ seen) of + result as {outcome = NONE, proof, used_thm_names, ...} : prover_result => + sublinear_minimize test (filter_used_facts used_thm_names xs) + (filter_used_facts used_thm_names seen, result) + | _ => sublinear_minimize test xs (x :: seen, result) + +(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer + preprocessing time is included in the estimate below but isn't part of the + timeout. *) +val fudge_msecs = 1000 + +fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set." + | minimize_theorems (params as {debug, atps = atp :: _, full_types, + isar_proof, isar_shrink_factor, timeout, ...}) + i n state axioms = + let + val thy = Proof.theory_of state + val prover = get_prover_fun thy atp + val msecs = Time.toMilliseconds timeout + val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".") + val {context = ctxt, goal, ...} = Proof.goal state + val (_, hyp_ts, concl_t) = strip_subgoal goal i + val explicit_apply = + not (forall (Meson.is_fol_term thy) + (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms)) + fun do_test timeout = + test_theorems params prover explicit_apply timeout i state + val timer = Timer.startRealTimer () + in + (case do_test timeout axioms of + result as {outcome = NONE, pool, used_thm_names, + conjecture_shape, ...} => + let + val time = Timer.checkRealTimer timer + val new_timeout = + Int.min (msecs, Time.toMilliseconds time + fudge_msecs) + |> Time.fromMilliseconds + val (min_thms, {proof, axiom_names, ...}) = + sublinear_minimize (do_test new_timeout) + (filter_used_facts used_thm_names axioms) ([], result) + val n = length min_thms + val _ = priority (cat_lines + ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^ + (case length (filter (curry (op =) Chained o snd o fst) min_thms) of + 0 => "" + | n => " (including " ^ Int.toString n ^ " chained)") ^ ".") + in + (SOME min_thms, + proof_text isar_proof + (pool, debug, isar_shrink_factor, ctxt, conjecture_shape) + (full_types, K "", proof, axiom_names, goal, i) |> fst) + end + | {outcome = SOME TimedOut, ...} => + (NONE, "Timeout: You can increase the time limit using the \"timeout\" \ + \option (e.g., \"timeout = " ^ + string_of_int (10 + msecs div 1000) ^ " s\").") + | {outcome = SOME UnknownError, ...} => + (* Failure sometimes mean timeout, unfortunately. *) + (NONE, "Failure: No proof was found with the current time limit. You \ + \can increase the time limit using the \"timeout\" \ + \option (e.g., \"timeout = " ^ + string_of_int (10 + msecs div 1000) ^ " s\").") + | {message, ...} => (NONE, "ATP error: " ^ message)) + handle ERROR msg => (NONE, "Error: " ^ msg) + end + +fun run_minimize params i refs state = + let + val ctxt = Proof.context_of state + val reserved = reserved_isar_keyword_table () + val chained_ths = #facts (Proof.goal state) + val axioms = + maps (map (apsnd single) + o name_thm_pairs_from_ref ctxt reserved chained_ths) refs + in + case subgoal_count state of + 0 => priority "No subgoal!" + | n => + (kill_atps (); priority (#2 (minimize_theorems params i n state axioms))) + end + +end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Thu Sep 02 08:40:25 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1038 +0,0 @@ -(* Title: HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML - Author: Lawrence C. Paulson, Cambridge University Computer Laboratory - Author: Claire Quigley, Cambridge University Computer Laboratory - Author: Jasmin Blanchette, TU Muenchen - -Transfer of proofs from external provers. -*) - -signature SLEDGEHAMMER_PROOF_RECONSTRUCT = -sig - type locality = Sledgehammer_Fact_Filter.locality - type minimize_command = string list -> string - type metis_params = - bool * minimize_command * string * (string * locality) list vector * thm - * int - type isar_params = - string Symtab.table * bool * int * Proof.context * int list list - type text_result = string * (string * locality) list - - val metis_proof_text : metis_params -> text_result - val isar_proof_text : isar_params -> metis_params -> text_result - val proof_text : bool -> isar_params -> metis_params -> text_result -end; - -structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT = -struct - -open ATP_Problem -open Metis_Clauses -open Sledgehammer_Util -open Sledgehammer_Fact_Filter -open Sledgehammer_Translate - -type minimize_command = string list -> string -type metis_params = - bool * minimize_command * string * (string * locality) list vector * thm * int -type isar_params = - string Symtab.table * bool * int * Proof.context * int list list -type text_result = string * (string * locality) list - -(* Simple simplifications to ensure that sort annotations don't leave a trail of - spurious "True"s. *) -fun s_not @{const False} = @{const True} - | s_not @{const True} = @{const False} - | s_not (@{const Not} $ t) = t - | s_not t = @{const Not} $ t -fun s_conj (@{const True}, t2) = t2 - | s_conj (t1, @{const True}) = t1 - | s_conj p = HOLogic.mk_conj p -fun s_disj (@{const False}, t2) = t2 - | s_disj (t1, @{const False}) = t1 - | s_disj p = HOLogic.mk_disj p -fun s_imp (@{const True}, t2) = t2 - | s_imp (t1, @{const False}) = s_not t1 - | s_imp p = HOLogic.mk_imp p -fun s_iff (@{const True}, t2) = t2 - | s_iff (t1, @{const True}) = t1 - | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2 - -fun mk_anot (AConn (ANot, [phi])) = phi - | mk_anot phi = AConn (ANot, [phi]) -fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2]) - -fun index_in_shape x = find_index (exists (curry (op =) x)) -fun is_axiom_number axiom_names num = - num > 0 andalso num <= Vector.length axiom_names andalso - not (null (Vector.sub (axiom_names, num - 1))) -fun is_conjecture_number conjecture_shape num = - index_in_shape num conjecture_shape >= 0 - -fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) = - Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t') - | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) = - Const (@{const_name All}, T) $ Abs (s, T', negate_term t') - | negate_term (@{const HOL.implies} $ t1 $ t2) = - @{const HOL.conj} $ t1 $ negate_term t2 - | negate_term (@{const HOL.conj} $ t1 $ t2) = - @{const HOL.disj} $ negate_term t1 $ negate_term t2 - | negate_term (@{const HOL.disj} $ t1 $ t2) = - @{const HOL.conj} $ negate_term t1 $ negate_term t2 - | negate_term (@{const Not} $ t) = t - | negate_term t = @{const Not} $ t - -datatype ('a, 'b, 'c, 'd, 'e) raw_step = - Definition of 'a * 'b * 'c | - Inference of 'a * 'd * 'e list - -fun raw_step_number (Definition (num, _, _)) = num - | raw_step_number (Inference (num, _, _)) = num - -(**** PARSING OF TSTP FORMAT ****) - -(*Strings enclosed in single quotes, e.g. filenames*) -val scan_quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode; - -val scan_dollar_name = - Scan.repeat ($$ "$") -- Symbol.scan_id >> (fn (ss, s) => implode ss ^ s) - -fun repair_name _ "$true" = "c_True" - | repair_name _ "$false" = "c_False" - | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *) - | repair_name _ "equal" = "c_equal" (* needed by SPASS? *) - | repair_name pool s = - case Symtab.lookup pool s of - SOME s' => s' - | NONE => - if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then - "c_equal" (* seen in Vampire proofs *) - else - s -(* Generalized first-order terms, which include file names, numbers, etc. *) -val parse_potential_integer = - (scan_dollar_name || scan_quoted) >> K NONE - || scan_integer >> SOME -fun parse_annotation x = - ((parse_potential_integer ::: Scan.repeat ($$ " " |-- parse_potential_integer) - >> map_filter I) -- Scan.optional parse_annotation [] - >> uncurry (union (op =)) - || $$ "(" |-- parse_annotations --| $$ ")" - || $$ "[" |-- parse_annotations --| $$ "]") x -and parse_annotations x = - (Scan.optional (parse_annotation - ::: Scan.repeat ($$ "," |-- parse_annotation)) [] - >> (fn numss => fold (union (op =)) numss [])) x - -(* Vampire proof lines sometimes contain needless information such as "(0:3)", - which can be hard to disambiguate from function application in an LL(1) - parser. As a workaround, we extend the TPTP term syntax with such detritus - and ignore it. *) -fun parse_vampire_detritus x = - (scan_integer |-- $$ ":" --| scan_integer >> K []) x - -fun parse_term pool x = - ((scan_dollar_name >> repair_name pool) - -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool) - --| $$ ")") [] - --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") [] - >> ATerm) x -and parse_terms pool x = - (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x - -fun parse_atom pool = - parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "=" - -- parse_term pool) - >> (fn (u1, NONE) => AAtom u1 - | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2])) - | (u1, SOME (SOME _, u2)) => - mk_anot (AAtom (ATerm ("c_equal", [u1, u2])))) - -fun fo_term_head (ATerm (s, _)) = s - -(* TPTP formulas are fully parenthesized, so we don't need to worry about - operator precedence. *) -fun parse_formula pool x = - (($$ "(" |-- parse_formula pool --| $$ ")" - || ($$ "!" >> K AForall || $$ "?" >> K AExists) - --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":" - -- parse_formula pool - >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi)) - || $$ "~" |-- parse_formula pool >> mk_anot - || parse_atom pool) - -- Scan.option ((Scan.this_string "=>" >> K AImplies - || Scan.this_string "<=>" >> K AIff - || Scan.this_string "<~>" >> K ANotIff - || Scan.this_string "<=" >> K AIf - || $$ "|" >> K AOr || $$ "&" >> K AAnd) - -- parse_formula pool) - >> (fn (phi1, NONE) => phi1 - | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x - -val parse_tstp_extra_arguments = - Scan.optional ($$ "," |-- parse_annotation - --| Scan.option ($$ "," |-- parse_annotations)) [] - -(* Syntax: (fof|cnf)\(, , \). - The could be an identifier, but we assume integers. *) - fun parse_tstp_line pool = - ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(") - |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ "," - -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "." - >> (fn (((num, role), phi), deps) => - case role of - "definition" => - (case phi of - AConn (AIff, [phi1 as AAtom _, phi2]) => - Definition (num, phi1, phi2) - | AAtom (ATerm ("c_equal", _)) => - Inference (num, phi, deps) (* Vampire's equality proxy axiom *) - | _ => raise Fail "malformed definition") - | _ => Inference (num, phi, deps)) - -(**** PARSING OF VAMPIRE OUTPUT ****) - -(* Syntax: . *) -fun parse_vampire_line pool = - scan_integer --| $$ "." -- parse_formula pool -- parse_annotation - >> (fn ((num, phi), deps) => Inference (num, phi, deps)) - -(**** PARSING OF SPASS OUTPUT ****) - -(* SPASS returns clause references of the form "x.y". We ignore "y", whose role - is not clear anyway. *) -val parse_dot_name = scan_integer --| $$ "." --| scan_integer - -val parse_spass_annotations = - Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name - --| Scan.option ($$ ","))) [] - -(* It is not clear why some literals are followed by sequences of stars and/or - pluses. We ignore them. *) -fun parse_decorated_atom pool = - parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ") - -fun mk_horn ([], []) = AAtom (ATerm ("c_False", [])) - | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits - | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits) - | mk_horn (neg_lits, pos_lits) = - mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits, - foldr1 (mk_aconn AOr) pos_lits) - -fun parse_horn_clause pool = - Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|" - -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">" - -- Scan.repeat (parse_decorated_atom pool) - >> (mk_horn o apfst (op @)) - -(* Syntax: [0:] - || -> . *) -fun parse_spass_line pool = - scan_integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id - -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "." - >> (fn ((num, deps), u) => Inference (num, u, deps)) - -fun parse_line pool = - parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool -fun parse_lines pool = Scan.repeat1 (parse_line pool) -fun parse_proof pool = - fst o Scan.finite Symbol.stopper - (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output") - (parse_lines pool))) - o explode o strip_spaces_except_between_ident_chars - -(**** INTERPRETATION OF TSTP SYNTAX TREES ****) - -exception FO_TERM of string fo_term list -exception FORMULA of (string, string fo_term) formula list -exception SAME of unit - -(* Type variables are given the basic sort "HOL.type". Some will later be - constrained by information from type literals, or by type inference. *) -fun type_from_fo_term tfrees (u as ATerm (a, us)) = - let val Ts = map (type_from_fo_term tfrees) us in - case strip_prefix_and_unascii type_const_prefix a of - SOME b => Type (invert_const b, Ts) - | NONE => - if not (null us) then - raise FO_TERM [u] (* only "tconst"s have type arguments *) - else case strip_prefix_and_unascii tfree_prefix a of - SOME b => - let val s = "'" ^ b in - TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS) - end - | NONE => - case strip_prefix_and_unascii tvar_prefix a of - SOME b => TVar (("'" ^ b, 0), HOLogic.typeS) - | NONE => - (* Variable from the ATP, say "X1" *) - Type_Infer.param 0 (a, HOLogic.typeS) - end - -(* Type class literal applied to a type. Returns triple of polarity, class, - type. *) -fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) = - case (strip_prefix_and_unascii class_prefix a, - map (type_from_fo_term tfrees) us) of - (SOME b, [T]) => (pos, b, T) - | _ => raise FO_TERM [u] - -(** Accumulate type constraints in a formula: negative type literals **) -fun add_var (key, z) = Vartab.map_default (key, []) (cons z) -fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl) - | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl) - | add_type_constraint _ = I - -fun repair_atp_variable_name f s = - let - fun subscript_name s n = s ^ nat_subscript n - val s = String.map f s - in - case space_explode "_" s of - [_] => (case take_suffix Char.isDigit (String.explode s) of - (cs1 as _ :: _, cs2 as _ :: _) => - subscript_name (String.implode cs1) - (the (Int.fromString (String.implode cs2))) - | (_, _) => s) - | [s1, s2] => (case Int.fromString s2 of - SOME n => subscript_name s1 n - | NONE => s) - | _ => s - end - -(* First-order translation. No types are known for variables. "HOLogic.typeT" - should allow them to be inferred. *) -fun raw_term_from_pred thy full_types tfrees = - let - fun aux opt_T extra_us u = - case u of - ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1 - | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1 - | ATerm (a, us) => - if a = type_wrapper_name then - case us of - [typ_u, term_u] => - aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u - | _ => raise FO_TERM us - else case strip_prefix_and_unascii const_prefix a of - SOME "equal" => - list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), - map (aux NONE []) us) - | SOME b => - let - val c = invert_const b - val num_type_args = num_type_args thy c - val (type_us, term_us) = - chop (if full_types then 0 else num_type_args) us - (* Extra args from "hAPP" come after any arguments given directly to - the constant. *) - val term_ts = map (aux NONE []) term_us - val extra_ts = map (aux NONE []) extra_us - val t = - Const (c, if full_types then - case opt_T of - SOME T => map fastype_of term_ts ---> T - | NONE => - if num_type_args = 0 then - Sign.const_instance thy (c, []) - else - raise Fail ("no type information for " ^ quote c) - else - Sign.const_instance thy (c, - map (type_from_fo_term tfrees) type_us)) - in list_comb (t, term_ts @ extra_ts) end - | NONE => (* a free or schematic variable *) - let - val ts = map (aux NONE []) (us @ extra_us) - val T = map fastype_of ts ---> HOLogic.typeT - val t = - case strip_prefix_and_unascii fixed_var_prefix a of - SOME b => Free (b, T) - | NONE => - case strip_prefix_and_unascii schematic_var_prefix a of - SOME b => Var ((b, 0), T) - | NONE => - if is_tptp_variable a then - Var ((repair_atp_variable_name Char.toLower a, 0), T) - else - (* Skolem constants? *) - Var ((repair_atp_variable_name Char.toUpper a, 0), T) - in list_comb (t, ts) end - in aux (SOME HOLogic.boolT) [] end - -fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) = - if String.isPrefix class_prefix s then - add_type_constraint (type_constraint_from_term pos tfrees u) - #> pair @{const True} - else - pair (raw_term_from_pred thy full_types tfrees u) - -val combinator_table = - [(@{const_name COMBI}, @{thm COMBI_def_raw}), - (@{const_name COMBK}, @{thm COMBK_def_raw}), - (@{const_name COMBB}, @{thm COMBB_def_raw}), - (@{const_name COMBC}, @{thm COMBC_def_raw}), - (@{const_name COMBS}, @{thm COMBS_def_raw})] - -fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2)) - | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t') - | uncombine_term (t as Const (x as (s, _))) = - (case AList.lookup (op =) combinator_table s of - SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd - | NONE => t) - | uncombine_term t = t - -(* Update schematic type variables with detected sort constraints. It's not - totally clear when this code is necessary. *) -fun repair_tvar_sorts (t, tvar_tab) = - let - fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) - | do_type (TVar (xi, s)) = - TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) - | do_type (TFree z) = TFree z - fun do_term (Const (a, T)) = Const (a, do_type T) - | do_term (Free (a, T)) = Free (a, do_type T) - | do_term (Var (xi, T)) = Var (xi, do_type T) - | do_term (t as Bound _) = t - | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) - | do_term (t1 $ t2) = do_term t1 $ do_term t2 - in t |> not (Vartab.is_empty tvar_tab) ? do_term end - -fun quantify_over_free quant_s free_s body_t = - case Term.add_frees body_t [] |> filter (curry (op =) free_s o fst) of - [] => body_t - | frees as (_, free_T) :: _ => - Abs (free_s, free_T, fold (curry abstract_over) (map Free frees) body_t) - -(* Interpret an ATP formula as a HOL term, extracting sort constraints as they - appear in the formula. *) -fun prop_from_formula thy full_types tfrees phi = - let - fun do_formula pos phi = - case phi of - AQuant (_, [], phi) => do_formula pos phi - | AQuant (q, x :: xs, phi') => - do_formula pos (AQuant (q, xs, phi')) - #>> quantify_over_free (case q of - AForall => @{const_name All} - | AExists => @{const_name Ex}) - (repair_atp_variable_name Char.toLower x) - | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not - | AConn (c, [phi1, phi2]) => - do_formula (pos |> c = AImplies ? not) phi1 - ##>> do_formula pos phi2 - #>> (case c of - AAnd => s_conj - | AOr => s_disj - | AImplies => s_imp - | AIf => s_imp o swap - | AIff => s_iff - | ANotIff => s_not o s_iff) - | AAtom tm => term_from_pred thy full_types tfrees pos tm - | _ => raise FORMULA [phi] - in repair_tvar_sorts (do_formula true phi Vartab.empty) end - -fun check_formula ctxt = - Type_Infer.constrain HOLogic.boolT - #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) - - -(**** Translation of TSTP files to Isar Proofs ****) - -fun unvarify_term (Var ((s, 0), T)) = Free (s, T) - | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t]) - -fun decode_line full_types tfrees (Definition (num, phi1, phi2)) ctxt = - let - val thy = ProofContext.theory_of ctxt - val t1 = prop_from_formula thy full_types tfrees phi1 - val vars = snd (strip_comb t1) - val frees = map unvarify_term vars - val unvarify_args = subst_atomic (vars ~~ frees) - val t2 = prop_from_formula thy full_types tfrees phi2 - val (t1, t2) = - HOLogic.eq_const HOLogic.typeT $ t1 $ t2 - |> unvarify_args |> uncombine_term |> check_formula ctxt - |> HOLogic.dest_eq - in - (Definition (num, t1, t2), - fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt) - end - | decode_line full_types tfrees (Inference (num, u, deps)) ctxt = - let - val thy = ProofContext.theory_of ctxt - val t = u |> prop_from_formula thy full_types tfrees - |> uncombine_term |> check_formula ctxt - in - (Inference (num, t, deps), - fold Variable.declare_term (OldTerm.term_frees t) ctxt) - end -fun decode_lines ctxt full_types tfrees lines = - fst (fold_map (decode_line full_types tfrees) lines ctxt) - -fun is_same_inference _ (Definition _) = false - | is_same_inference t (Inference (_, t', _)) = t aconv t' - -(* No "real" literals means only type information (tfree_tcs, clsrel, or - clsarity). *) -val is_only_type_information = curry (op aconv) HOLogic.true_const - -fun replace_one_dep (old, new) dep = if dep = old then new else [dep] -fun replace_deps_in_line _ (line as Definition _) = line - | replace_deps_in_line p (Inference (num, t, deps)) = - Inference (num, t, fold (union (op =) o replace_one_dep p) deps []) - -(* Discard axioms; consolidate adjacent lines that prove the same formula, since - they differ only in type information.*) -fun add_line _ _ (line as Definition _) lines = line :: lines - | add_line conjecture_shape axiom_names (Inference (num, t, [])) lines = - (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or - definitions. *) - if is_axiom_number axiom_names num then - (* Axioms are not proof lines. *) - if is_only_type_information t then - map (replace_deps_in_line (num, [])) lines - (* Is there a repetition? If so, replace later line by earlier one. *) - else case take_prefix (not o is_same_inference t) lines of - (_, []) => lines (*no repetition of proof line*) - | (pre, Inference (num', _, _) :: post) => - pre @ map (replace_deps_in_line (num', [num])) post - else if is_conjecture_number conjecture_shape num then - Inference (num, negate_term t, []) :: lines - else - map (replace_deps_in_line (num, [])) lines - | add_line _ _ (Inference (num, t, deps)) lines = - (* Type information will be deleted later; skip repetition test. *) - if is_only_type_information t then - Inference (num, t, deps) :: lines - (* Is there a repetition? If so, replace later line by earlier one. *) - else case take_prefix (not o is_same_inference t) lines of - (* FIXME: Doesn't this code risk conflating proofs involving different - types? *) - (_, []) => Inference (num, t, deps) :: lines - | (pre, Inference (num', t', _) :: post) => - Inference (num, t', deps) :: - pre @ map (replace_deps_in_line (num', [num])) post - -(* Recursively delete empty lines (type information) from the proof. *) -fun add_nontrivial_line (Inference (num, t, [])) lines = - if is_only_type_information t then delete_dep num lines - else Inference (num, t, []) :: lines - | add_nontrivial_line line lines = line :: lines -and delete_dep num lines = - fold_rev add_nontrivial_line (map (replace_deps_in_line (num, [])) lines) [] - -(* ATPs sometimes reuse free variable names in the strangest ways. Removing - offending lines often does the trick. *) -fun is_bad_free frees (Free x) = not (member (op =) frees x) - | is_bad_free _ _ = false - -(* Vampire is keen on producing these. *) -fun is_trivial_formula (@{const Not} $ (Const (@{const_name HOL.eq}, _) - $ t1 $ t2)) = (t1 aconv t2) - | is_trivial_formula _ = false - -fun add_desired_line _ _ _ _ (line as Definition (num, _, _)) (j, lines) = - (j, line :: map (replace_deps_in_line (num, [])) lines) - | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees - (Inference (num, t, deps)) (j, lines) = - (j + 1, - if is_axiom_number axiom_names num orelse - is_conjecture_number conjecture_shape num orelse - (not (is_only_type_information t) andalso - null (Term.add_tvars t []) andalso - not (exists_subterm (is_bad_free frees) t) andalso - not (is_trivial_formula t) andalso - (null lines orelse (* last line must be kept *) - (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then - Inference (num, t, deps) :: lines (* keep line *) - else - map (replace_deps_in_line (num, deps)) lines) (* drop line *) - -(** EXTRACTING LEMMAS **) - -(* Like "split_line", but ignores "\n" that follow a comma (as in SNARK's - output). *) -val split_proof_lines = - let - fun aux [] [] = [] - | aux line [] = [implode (rev line)] - | aux line ("," :: "\n" :: rest) = aux ("," :: line) rest - | aux line ("\n" :: rest) = aux line [] @ aux [] rest - | aux line (s :: rest) = aux (s :: line) rest - in aux [] o explode end - -(* A list consisting of the first number in each line is returned. For TSTP, - interesting lines have the form "fof(108, axiom, ...)", where the number - (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where - the first number (108) is extracted. For Vampire, we look for - "108. ... [input]". *) -fun used_facts_in_atp_proof axiom_names atp_proof = - let - fun axiom_names_at_index num = - let val j = Int.fromString num |> the_default ~1 in - if is_axiom_number axiom_names j then Vector.sub (axiom_names, j - 1) - else [] - end - val tokens_of = - String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_") - fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) = - if tag = "cnf" orelse tag = "fof" then - (case strip_prefix_and_unascii axiom_prefix (List.last rest) of - SOME name => - if member (op =) rest "file" then - ([(name, name |> find_first_in_list_vector axiom_names |> the)] - handle Option.Option => - error ("No such fact: " ^ quote name ^ ".")) - else - axiom_names_at_index num - | NONE => axiom_names_at_index num) - else - [] - | do_line (num :: "0" :: "Inp" :: _) = axiom_names_at_index num - | do_line (num :: rest) = - (case List.last rest of "input" => axiom_names_at_index num | _ => []) - | do_line _ = [] - in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end - -val indent_size = 2 -val no_label = ("", ~1) - -val raw_prefix = "X" -val assum_prefix = "A" -val fact_prefix = "F" - -fun string_for_label (s, num) = s ^ string_of_int num - -fun metis_using [] = "" - | metis_using ls = - "using " ^ space_implode " " (map string_for_label ls) ^ " " -fun metis_apply _ 1 = "by " - | metis_apply 1 _ = "apply " - | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply " -fun metis_name full_types = if full_types then "metisFT" else "metis" -fun metis_call full_types [] = metis_name full_types - | metis_call full_types ss = - "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")" -fun metis_command full_types i n (ls, ss) = - metis_using ls ^ metis_apply i n ^ metis_call full_types ss -fun metis_line full_types i n ss = - "Try this command: " ^ - Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "." -fun minimize_line _ [] = "" - | minimize_line minimize_command ss = - case minimize_command ss of - "" => "" - | command => - "\nTo minimize the number of lemmas, try this: " ^ - Markup.markup Markup.sendback command ^ "." - -fun used_facts axiom_names = - used_facts_in_atp_proof axiom_names - #> List.partition (curry (op =) Chained o snd) - #> pairself (sort_distinct (string_ord o pairself fst)) - -fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names, - goal, i) = - let - val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof - val n = Logic.count_prems (prop_of goal) - in - (metis_line full_types i n (map fst other_lemmas) ^ - minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)), - other_lemmas @ chained_lemmas) - end - -(** Isar proof construction and manipulation **) - -fun merge_fact_sets (ls1, ss1) (ls2, ss2) = - (union (op =) ls1 ls2, union (op =) ss1 ss2) - -type label = string * int -type facts = label list * string list - -datatype qualifier = Show | Then | Moreover | Ultimately - -datatype step = - Fix of (string * typ) list | - Let of term * term | - Assume of label * term | - Have of qualifier list * label * term * byline -and byline = - ByMetis of facts | - CaseSplit of step list list * facts - -fun smart_case_split [] facts = ByMetis facts - | smart_case_split proofs facts = CaseSplit (proofs, facts) - -fun add_fact_from_dep axiom_names num = - if is_axiom_number axiom_names num then - apsnd (union (op =) (map fst (Vector.sub (axiom_names, num - 1)))) - else - apfst (insert (op =) (raw_prefix, num)) - -fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t -fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t - -fun step_for_line _ _ (Definition (_, t1, t2)) = Let (t1, t2) - | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t) - | step_for_line axiom_names j (Inference (num, t, deps)) = - Have (if j = 1 then [Show] else [], (raw_prefix, num), - forall_vars t, - ByMetis (fold (add_fact_from_dep axiom_names) deps ([], []))) - -fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor - atp_proof conjecture_shape axiom_names params frees = - let - val lines = - atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *) - |> parse_proof pool - |> sort (int_ord o pairself raw_step_number) - |> decode_lines ctxt full_types tfrees - |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names) - |> rpair [] |-> fold_rev add_nontrivial_line - |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor - conjecture_shape axiom_names frees) - |> snd - in - (if null params then [] else [Fix params]) @ - map2 (step_for_line axiom_names) (length lines downto 1) lines - end - -(* When redirecting proofs, we keep information about the labels seen so far in - the "backpatches" data structure. The first component indicates which facts - should be associated with forthcoming proof steps. The second component is a - pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should - become assumptions and "drop_ls" are the labels that should be dropped in a - case split. *) -type backpatches = (label * facts) list * (label list * label list) - -fun used_labels_of_step (Have (_, _, _, by)) = - (case by of - ByMetis (ls, _) => ls - | CaseSplit (proofs, (ls, _)) => - fold (union (op =) o used_labels_of) proofs ls) - | used_labels_of_step _ = [] -and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof [] - -fun new_labels_of_step (Fix _) = [] - | new_labels_of_step (Let _) = [] - | new_labels_of_step (Assume (l, _)) = [l] - | new_labels_of_step (Have (_, l, _, _)) = [l] -val new_labels_of = maps new_labels_of_step - -val join_proofs = - let - fun aux _ [] = NONE - | aux proof_tail (proofs as (proof1 :: _)) = - if exists null proofs then - NONE - else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then - aux (hd proof1 :: proof_tail) (map tl proofs) - else case hd proof1 of - Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *) - if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t') - | _ => false) (tl proofs) andalso - not (exists (member (op =) (maps new_labels_of proofs)) - (used_labels_of proof_tail)) then - SOME (l, t, map rev proofs, proof_tail) - else - NONE - | _ => NONE - in aux [] o map rev end - -fun case_split_qualifiers proofs = - case length proofs of - 0 => [] - | 1 => [Then] - | _ => [Ultimately] - -fun redirect_proof conjecture_shape hyp_ts concl_t proof = - let - (* The first pass outputs those steps that are independent of the negated - conjecture. The second pass flips the proof by contradiction to obtain a - direct proof, introducing case splits when an inference depends on - several facts that depend on the negated conjecture. *) - fun find_hyp num = - nth hyp_ts (index_in_shape num conjecture_shape) - handle Subscript => - raise Fail ("Cannot find hypothesis " ^ Int.toString num) - val concl_ls = map (pair raw_prefix) (List.last conjecture_shape) - val canonicalize_labels = - map (fn l => if member (op =) concl_ls l then hd concl_ls else l) - #> distinct (op =) - fun first_pass ([], contra) = ([], contra) - | first_pass ((step as Fix _) :: proof, contra) = - first_pass (proof, contra) |>> cons step - | first_pass ((step as Let _) :: proof, contra) = - first_pass (proof, contra) |>> cons step - | first_pass ((step as Assume (l as (_, num), _)) :: proof, contra) = - if member (op =) concl_ls l then - first_pass (proof, contra ||> l = hd concl_ls ? cons step) - else - first_pass (proof, contra) |>> cons (Assume (l, find_hyp num)) - | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) = - let - val ls = canonicalize_labels ls - val step = Have (qs, l, t, ByMetis (ls, ss)) - in - if exists (member (op =) (fst contra)) ls then - first_pass (proof, contra |>> cons l ||> cons step) - else - first_pass (proof, contra) |>> cons step - end - | first_pass _ = raise Fail "malformed proof" - val (proof_top, (contra_ls, contra_proof)) = - first_pass (proof, (concl_ls, [])) - val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst - fun backpatch_labels patches ls = - fold merge_fact_sets (map (backpatch_label patches) ls) ([], []) - fun second_pass end_qs ([], assums, patches) = - ([Have (end_qs, no_label, concl_t, - ByMetis (backpatch_labels patches (map snd assums)))], patches) - | second_pass end_qs (Assume (l, t) :: proof, assums, patches) = - second_pass end_qs (proof, (t, l) :: assums, patches) - | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums, - patches) = - if member (op =) (snd (snd patches)) l andalso - not (member (op =) (fst (snd patches)) l) andalso - not (AList.defined (op =) (fst patches) l) then - second_pass end_qs (proof, assums, patches ||> apsnd (append ls)) - else - (case List.partition (member (op =) contra_ls) ls of - ([contra_l], co_ls) => - if member (op =) qs Show then - second_pass end_qs (proof, assums, - patches |>> cons (contra_l, (co_ls, ss))) - else - second_pass end_qs - (proof, assums, - patches |>> cons (contra_l, (l :: co_ls, ss))) - |>> cons (if member (op =) (fst (snd patches)) l then - Assume (l, negate_term t) - else - Have (qs, l, negate_term t, - ByMetis (backpatch_label patches l))) - | (contra_ls as _ :: _, co_ls) => - let - val proofs = - map_filter - (fn l => - if member (op =) concl_ls l then - NONE - else - let - val drop_ls = filter (curry (op <>) l) contra_ls - in - second_pass [] - (proof, assums, - patches ||> apfst (insert (op =) l) - ||> apsnd (union (op =) drop_ls)) - |> fst |> SOME - end) contra_ls - val (assumes, facts) = - if member (op =) (fst (snd patches)) l then - ([Assume (l, negate_term t)], (l :: co_ls, ss)) - else - ([], (co_ls, ss)) - in - (case join_proofs proofs of - SOME (l, t, proofs, proof_tail) => - Have (case_split_qualifiers proofs @ - (if null proof_tail then end_qs else []), l, t, - smart_case_split proofs facts) :: proof_tail - | NONE => - [Have (case_split_qualifiers proofs @ end_qs, no_label, - concl_t, smart_case_split proofs facts)], - patches) - |>> append assumes - end - | _ => raise Fail "malformed proof") - | second_pass _ _ = raise Fail "malformed proof" - val proof_bottom = - second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst - in proof_top @ proof_bottom end - -(* FIXME: Still needed? Probably not. *) -val kill_duplicate_assumptions_in_proof = - let - fun relabel_facts subst = - apfst (map (fn l => AList.lookup (op =) subst l |> the_default l)) - fun do_step (step as Assume (l, t)) (proof, subst, assums) = - (case AList.lookup (op aconv) assums t of - SOME l' => (proof, (l, l') :: subst, assums) - | NONE => (step :: proof, subst, (t, l) :: assums)) - | do_step (Have (qs, l, t, by)) (proof, subst, assums) = - (Have (qs, l, t, - case by of - ByMetis facts => ByMetis (relabel_facts subst facts) - | CaseSplit (proofs, facts) => - CaseSplit (map do_proof proofs, relabel_facts subst facts)) :: - proof, subst, assums) - | do_step step (proof, subst, assums) = (step :: proof, subst, assums) - and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev - in do_proof end - -val then_chain_proof = - let - fun aux _ [] = [] - | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof - | aux l' (Have (qs, l, t, by) :: proof) = - (case by of - ByMetis (ls, ss) => - Have (if member (op =) ls l' then - (Then :: qs, l, t, - ByMetis (filter_out (curry (op =) l') ls, ss)) - else - (qs, l, t, ByMetis (ls, ss))) - | CaseSplit (proofs, facts) => - Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) :: - aux l proof - | aux _ (step :: proof) = step :: aux no_label proof - in aux no_label end - -fun kill_useless_labels_in_proof proof = - let - val used_ls = used_labels_of proof - fun do_label l = if member (op =) used_ls l then l else no_label - fun do_step (Assume (l, t)) = Assume (do_label l, t) - | do_step (Have (qs, l, t, by)) = - Have (qs, do_label l, t, - case by of - CaseSplit (proofs, facts) => - CaseSplit (map (map do_step) proofs, facts) - | _ => by) - | do_step step = step - in map do_step proof end - -fun prefix_for_depth n = replicate_string (n + 1) - -val relabel_proof = - let - fun aux _ _ _ [] = [] - | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) = - if l = no_label then - Assume (l, t) :: aux subst depth (next_assum, next_fact) proof - else - let val l' = (prefix_for_depth depth assum_prefix, next_assum) in - Assume (l', t) :: - aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof - end - | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) = - let - val (l', subst, next_fact) = - if l = no_label then - (l, subst, next_fact) - else - let - val l' = (prefix_for_depth depth fact_prefix, next_fact) - in (l', (l, l') :: subst, next_fact + 1) end - val relabel_facts = - apfst (map (fn l => - case AList.lookup (op =) subst l of - SOME l' => l' - | NONE => raise Fail ("unknown label " ^ - quote (string_for_label l)))) - val by = - case by of - ByMetis facts => ByMetis (relabel_facts facts) - | CaseSplit (proofs, facts) => - CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs, - relabel_facts facts) - in - Have (qs, l', t, by) :: - aux subst depth (next_assum, next_fact) proof - end - | aux subst depth nextp (step :: proof) = - step :: aux subst depth nextp proof - in aux [] 0 (1, 1) end - -fun string_for_proof ctxt full_types i n = - let - fun fix_print_mode f x = - setmp_CRITICAL show_no_free_types true - (setmp_CRITICAL show_types true - (Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) - (print_mode_value ())) f)) x - fun do_indent ind = replicate_string (ind * indent_size) " " - fun do_free (s, T) = - maybe_quote s ^ " :: " ^ - maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T) - fun do_label l = if l = no_label then "" else string_for_label l ^ ": " - fun do_have qs = - (if member (op =) qs Moreover then "moreover " else "") ^ - (if member (op =) qs Ultimately then "ultimately " else "") ^ - (if member (op =) qs Then then - if member (op =) qs Show then "thus" else "hence" - else - if member (op =) qs Show then "show" else "have") - val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt) - fun do_facts (ls, ss) = - metis_command full_types 1 1 - (ls |> sort_distinct (prod_ord string_ord int_ord), - ss |> sort_distinct string_ord) - and do_step ind (Fix xs) = - do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n" - | do_step ind (Let (t1, t2)) = - do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n" - | do_step ind (Assume (l, t)) = - do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n" - | do_step ind (Have (qs, l, t, ByMetis facts)) = - do_indent ind ^ do_have qs ^ " " ^ - do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n" - | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) = - space_implode (do_indent ind ^ "moreover\n") - (map (do_block ind) proofs) ^ - do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^ - do_facts facts ^ "\n" - and do_steps prefix suffix ind steps = - let val s = implode (map (do_step ind) steps) in - replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ - String.extract (s, ind * indent_size, - SOME (size s - ind * indent_size - 1)) ^ - suffix ^ "\n" - end - and do_block ind proof = do_steps "{ " " }" (ind + 1) proof - (* One-step proofs are pointless; better use the Metis one-liner - directly. *) - and do_proof [Have (_, _, _, ByMetis _)] = "" - | do_proof proof = - (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ - do_indent 0 ^ "proof -\n" ^ - do_steps "" "" 1 proof ^ - do_indent 0 ^ (if n <> 1 then "next" else "qed") - in do_proof end - -fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape) - (other_params as (full_types, _, atp_proof, axiom_names, - goal, i)) = - let - val (params, hyp_ts, concl_t) = strip_subgoal goal i - val frees = fold Term.add_frees (concl_t :: hyp_ts) [] - val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) [] - val n = Logic.count_prems (prop_of goal) - val (one_line_proof, lemma_names) = metis_proof_text other_params - fun isar_proof_for () = - case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor - atp_proof conjecture_shape axiom_names params - frees - |> redirect_proof conjecture_shape hyp_ts concl_t - |> kill_duplicate_assumptions_in_proof - |> then_chain_proof - |> kill_useless_labels_in_proof - |> relabel_proof - |> string_for_proof ctxt full_types i n of - "" => "\nNo structured proof available." - | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof - val isar_proof = - if debug then - isar_proof_for () - else - try isar_proof_for () - |> the_default "\nWarning: The Isar proof construction failed." - in (one_line_proof ^ isar_proof, lemma_names) end - -fun proof_text isar_proof isar_params other_params = - (if isar_proof then isar_proof_text isar_params else metis_proof_text) - other_params - -end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Thu Sep 02 08:51:16 2010 +0200 @@ -0,0 +1,1038 @@ +(* Title: HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML + Author: Lawrence C. Paulson, Cambridge University Computer Laboratory + Author: Claire Quigley, Cambridge University Computer Laboratory + Author: Jasmin Blanchette, TU Muenchen + +Transfer of proofs from external provers. +*) + +signature SLEDGEHAMMER_RECONSTRUCT = +sig + type locality = Sledgehammer_Filter.locality + type minimize_command = string list -> string + type metis_params = + bool * minimize_command * string * (string * locality) list vector * thm + * int + type isar_params = + string Symtab.table * bool * int * Proof.context * int list list + type text_result = string * (string * locality) list + + val metis_proof_text : metis_params -> text_result + val isar_proof_text : isar_params -> metis_params -> text_result + val proof_text : bool -> isar_params -> metis_params -> text_result +end; + +structure Sledgehammer_Reconstruct : SLEDGEHAMMER_RECONSTRUCT = +struct + +open ATP_Problem +open Metis_Clauses +open Sledgehammer_Util +open Sledgehammer_Filter +open Sledgehammer_Translate + +type minimize_command = string list -> string +type metis_params = + bool * minimize_command * string * (string * locality) list vector * thm * int +type isar_params = + string Symtab.table * bool * int * Proof.context * int list list +type text_result = string * (string * locality) list + +(* Simple simplifications to ensure that sort annotations don't leave a trail of + spurious "True"s. *) +fun s_not @{const False} = @{const True} + | s_not @{const True} = @{const False} + | s_not (@{const Not} $ t) = t + | s_not t = @{const Not} $ t +fun s_conj (@{const True}, t2) = t2 + | s_conj (t1, @{const True}) = t1 + | s_conj p = HOLogic.mk_conj p +fun s_disj (@{const False}, t2) = t2 + | s_disj (t1, @{const False}) = t1 + | s_disj p = HOLogic.mk_disj p +fun s_imp (@{const True}, t2) = t2 + | s_imp (t1, @{const False}) = s_not t1 + | s_imp p = HOLogic.mk_imp p +fun s_iff (@{const True}, t2) = t2 + | s_iff (t1, @{const True}) = t1 + | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2 + +fun mk_anot (AConn (ANot, [phi])) = phi + | mk_anot phi = AConn (ANot, [phi]) +fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2]) + +fun index_in_shape x = find_index (exists (curry (op =) x)) +fun is_axiom_number axiom_names num = + num > 0 andalso num <= Vector.length axiom_names andalso + not (null (Vector.sub (axiom_names, num - 1))) +fun is_conjecture_number conjecture_shape num = + index_in_shape num conjecture_shape >= 0 + +fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) = + Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t') + | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) = + Const (@{const_name All}, T) $ Abs (s, T', negate_term t') + | negate_term (@{const HOL.implies} $ t1 $ t2) = + @{const HOL.conj} $ t1 $ negate_term t2 + | negate_term (@{const HOL.conj} $ t1 $ t2) = + @{const HOL.disj} $ negate_term t1 $ negate_term t2 + | negate_term (@{const HOL.disj} $ t1 $ t2) = + @{const HOL.conj} $ negate_term t1 $ negate_term t2 + | negate_term (@{const Not} $ t) = t + | negate_term t = @{const Not} $ t + +datatype ('a, 'b, 'c, 'd, 'e) raw_step = + Definition of 'a * 'b * 'c | + Inference of 'a * 'd * 'e list + +fun raw_step_number (Definition (num, _, _)) = num + | raw_step_number (Inference (num, _, _)) = num + +(**** PARSING OF TSTP FORMAT ****) + +(*Strings enclosed in single quotes, e.g. filenames*) +val scan_quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode; + +val scan_dollar_name = + Scan.repeat ($$ "$") -- Symbol.scan_id >> (fn (ss, s) => implode ss ^ s) + +fun repair_name _ "$true" = "c_True" + | repair_name _ "$false" = "c_False" + | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *) + | repair_name _ "equal" = "c_equal" (* needed by SPASS? *) + | repair_name pool s = + case Symtab.lookup pool s of + SOME s' => s' + | NONE => + if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then + "c_equal" (* seen in Vampire proofs *) + else + s +(* Generalized first-order terms, which include file names, numbers, etc. *) +val parse_potential_integer = + (scan_dollar_name || scan_quoted) >> K NONE + || scan_integer >> SOME +fun parse_annotation x = + ((parse_potential_integer ::: Scan.repeat ($$ " " |-- parse_potential_integer) + >> map_filter I) -- Scan.optional parse_annotation [] + >> uncurry (union (op =)) + || $$ "(" |-- parse_annotations --| $$ ")" + || $$ "[" |-- parse_annotations --| $$ "]") x +and parse_annotations x = + (Scan.optional (parse_annotation + ::: Scan.repeat ($$ "," |-- parse_annotation)) [] + >> (fn numss => fold (union (op =)) numss [])) x + +(* Vampire proof lines sometimes contain needless information such as "(0:3)", + which can be hard to disambiguate from function application in an LL(1) + parser. As a workaround, we extend the TPTP term syntax with such detritus + and ignore it. *) +fun parse_vampire_detritus x = + (scan_integer |-- $$ ":" --| scan_integer >> K []) x + +fun parse_term pool x = + ((scan_dollar_name >> repair_name pool) + -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool) + --| $$ ")") [] + --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") [] + >> ATerm) x +and parse_terms pool x = + (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x + +fun parse_atom pool = + parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "=" + -- parse_term pool) + >> (fn (u1, NONE) => AAtom u1 + | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2])) + | (u1, SOME (SOME _, u2)) => + mk_anot (AAtom (ATerm ("c_equal", [u1, u2])))) + +fun fo_term_head (ATerm (s, _)) = s + +(* TPTP formulas are fully parenthesized, so we don't need to worry about + operator precedence. *) +fun parse_formula pool x = + (($$ "(" |-- parse_formula pool --| $$ ")" + || ($$ "!" >> K AForall || $$ "?" >> K AExists) + --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":" + -- parse_formula pool + >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi)) + || $$ "~" |-- parse_formula pool >> mk_anot + || parse_atom pool) + -- Scan.option ((Scan.this_string "=>" >> K AImplies + || Scan.this_string "<=>" >> K AIff + || Scan.this_string "<~>" >> K ANotIff + || Scan.this_string "<=" >> K AIf + || $$ "|" >> K AOr || $$ "&" >> K AAnd) + -- parse_formula pool) + >> (fn (phi1, NONE) => phi1 + | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x + +val parse_tstp_extra_arguments = + Scan.optional ($$ "," |-- parse_annotation + --| Scan.option ($$ "," |-- parse_annotations)) [] + +(* Syntax: (fof|cnf)\(, , \). + The could be an identifier, but we assume integers. *) + fun parse_tstp_line pool = + ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(") + |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ "," + -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "." + >> (fn (((num, role), phi), deps) => + case role of + "definition" => + (case phi of + AConn (AIff, [phi1 as AAtom _, phi2]) => + Definition (num, phi1, phi2) + | AAtom (ATerm ("c_equal", _)) => + Inference (num, phi, deps) (* Vampire's equality proxy axiom *) + | _ => raise Fail "malformed definition") + | _ => Inference (num, phi, deps)) + +(**** PARSING OF VAMPIRE OUTPUT ****) + +(* Syntax: . *) +fun parse_vampire_line pool = + scan_integer --| $$ "." -- parse_formula pool -- parse_annotation + >> (fn ((num, phi), deps) => Inference (num, phi, deps)) + +(**** PARSING OF SPASS OUTPUT ****) + +(* SPASS returns clause references of the form "x.y". We ignore "y", whose role + is not clear anyway. *) +val parse_dot_name = scan_integer --| $$ "." --| scan_integer + +val parse_spass_annotations = + Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name + --| Scan.option ($$ ","))) [] + +(* It is not clear why some literals are followed by sequences of stars and/or + pluses. We ignore them. *) +fun parse_decorated_atom pool = + parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ") + +fun mk_horn ([], []) = AAtom (ATerm ("c_False", [])) + | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits + | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits) + | mk_horn (neg_lits, pos_lits) = + mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits, + foldr1 (mk_aconn AOr) pos_lits) + +fun parse_horn_clause pool = + Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|" + -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">" + -- Scan.repeat (parse_decorated_atom pool) + >> (mk_horn o apfst (op @)) + +(* Syntax: [0:] + || -> . *) +fun parse_spass_line pool = + scan_integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id + -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "." + >> (fn ((num, deps), u) => Inference (num, u, deps)) + +fun parse_line pool = + parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool +fun parse_lines pool = Scan.repeat1 (parse_line pool) +fun parse_proof pool = + fst o Scan.finite Symbol.stopper + (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output") + (parse_lines pool))) + o explode o strip_spaces_except_between_ident_chars + +(**** INTERPRETATION OF TSTP SYNTAX TREES ****) + +exception FO_TERM of string fo_term list +exception FORMULA of (string, string fo_term) formula list +exception SAME of unit + +(* Type variables are given the basic sort "HOL.type". Some will later be + constrained by information from type literals, or by type inference. *) +fun type_from_fo_term tfrees (u as ATerm (a, us)) = + let val Ts = map (type_from_fo_term tfrees) us in + case strip_prefix_and_unascii type_const_prefix a of + SOME b => Type (invert_const b, Ts) + | NONE => + if not (null us) then + raise FO_TERM [u] (* only "tconst"s have type arguments *) + else case strip_prefix_and_unascii tfree_prefix a of + SOME b => + let val s = "'" ^ b in + TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS) + end + | NONE => + case strip_prefix_and_unascii tvar_prefix a of + SOME b => TVar (("'" ^ b, 0), HOLogic.typeS) + | NONE => + (* Variable from the ATP, say "X1" *) + Type_Infer.param 0 (a, HOLogic.typeS) + end + +(* Type class literal applied to a type. Returns triple of polarity, class, + type. *) +fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) = + case (strip_prefix_and_unascii class_prefix a, + map (type_from_fo_term tfrees) us) of + (SOME b, [T]) => (pos, b, T) + | _ => raise FO_TERM [u] + +(** Accumulate type constraints in a formula: negative type literals **) +fun add_var (key, z) = Vartab.map_default (key, []) (cons z) +fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl) + | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl) + | add_type_constraint _ = I + +fun repair_atp_variable_name f s = + let + fun subscript_name s n = s ^ nat_subscript n + val s = String.map f s + in + case space_explode "_" s of + [_] => (case take_suffix Char.isDigit (String.explode s) of + (cs1 as _ :: _, cs2 as _ :: _) => + subscript_name (String.implode cs1) + (the (Int.fromString (String.implode cs2))) + | (_, _) => s) + | [s1, s2] => (case Int.fromString s2 of + SOME n => subscript_name s1 n + | NONE => s) + | _ => s + end + +(* First-order translation. No types are known for variables. "HOLogic.typeT" + should allow them to be inferred. *) +fun raw_term_from_pred thy full_types tfrees = + let + fun aux opt_T extra_us u = + case u of + ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1 + | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1 + | ATerm (a, us) => + if a = type_wrapper_name then + case us of + [typ_u, term_u] => + aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u + | _ => raise FO_TERM us + else case strip_prefix_and_unascii const_prefix a of + SOME "equal" => + list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), + map (aux NONE []) us) + | SOME b => + let + val c = invert_const b + val num_type_args = num_type_args thy c + val (type_us, term_us) = + chop (if full_types then 0 else num_type_args) us + (* Extra args from "hAPP" come after any arguments given directly to + the constant. *) + val term_ts = map (aux NONE []) term_us + val extra_ts = map (aux NONE []) extra_us + val t = + Const (c, if full_types then + case opt_T of + SOME T => map fastype_of term_ts ---> T + | NONE => + if num_type_args = 0 then + Sign.const_instance thy (c, []) + else + raise Fail ("no type information for " ^ quote c) + else + Sign.const_instance thy (c, + map (type_from_fo_term tfrees) type_us)) + in list_comb (t, term_ts @ extra_ts) end + | NONE => (* a free or schematic variable *) + let + val ts = map (aux NONE []) (us @ extra_us) + val T = map fastype_of ts ---> HOLogic.typeT + val t = + case strip_prefix_and_unascii fixed_var_prefix a of + SOME b => Free (b, T) + | NONE => + case strip_prefix_and_unascii schematic_var_prefix a of + SOME b => Var ((b, 0), T) + | NONE => + if is_tptp_variable a then + Var ((repair_atp_variable_name Char.toLower a, 0), T) + else + (* Skolem constants? *) + Var ((repair_atp_variable_name Char.toUpper a, 0), T) + in list_comb (t, ts) end + in aux (SOME HOLogic.boolT) [] end + +fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) = + if String.isPrefix class_prefix s then + add_type_constraint (type_constraint_from_term pos tfrees u) + #> pair @{const True} + else + pair (raw_term_from_pred thy full_types tfrees u) + +val combinator_table = + [(@{const_name COMBI}, @{thm COMBI_def_raw}), + (@{const_name COMBK}, @{thm COMBK_def_raw}), + (@{const_name COMBB}, @{thm COMBB_def_raw}), + (@{const_name COMBC}, @{thm COMBC_def_raw}), + (@{const_name COMBS}, @{thm COMBS_def_raw})] + +fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2)) + | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t') + | uncombine_term (t as Const (x as (s, _))) = + (case AList.lookup (op =) combinator_table s of + SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd + | NONE => t) + | uncombine_term t = t + +(* Update schematic type variables with detected sort constraints. It's not + totally clear when this code is necessary. *) +fun repair_tvar_sorts (t, tvar_tab) = + let + fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) + | do_type (TVar (xi, s)) = + TVar (xi, the_default s (Vartab.lookup tvar_tab xi)) + | do_type (TFree z) = TFree z + fun do_term (Const (a, T)) = Const (a, do_type T) + | do_term (Free (a, T)) = Free (a, do_type T) + | do_term (Var (xi, T)) = Var (xi, do_type T) + | do_term (t as Bound _) = t + | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) + | do_term (t1 $ t2) = do_term t1 $ do_term t2 + in t |> not (Vartab.is_empty tvar_tab) ? do_term end + +fun quantify_over_free quant_s free_s body_t = + case Term.add_frees body_t [] |> filter (curry (op =) free_s o fst) of + [] => body_t + | frees as (_, free_T) :: _ => + Abs (free_s, free_T, fold (curry abstract_over) (map Free frees) body_t) + +(* Interpret an ATP formula as a HOL term, extracting sort constraints as they + appear in the formula. *) +fun prop_from_formula thy full_types tfrees phi = + let + fun do_formula pos phi = + case phi of + AQuant (_, [], phi) => do_formula pos phi + | AQuant (q, x :: xs, phi') => + do_formula pos (AQuant (q, xs, phi')) + #>> quantify_over_free (case q of + AForall => @{const_name All} + | AExists => @{const_name Ex}) + (repair_atp_variable_name Char.toLower x) + | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not + | AConn (c, [phi1, phi2]) => + do_formula (pos |> c = AImplies ? not) phi1 + ##>> do_formula pos phi2 + #>> (case c of + AAnd => s_conj + | AOr => s_disj + | AImplies => s_imp + | AIf => s_imp o swap + | AIff => s_iff + | ANotIff => s_not o s_iff) + | AAtom tm => term_from_pred thy full_types tfrees pos tm + | _ => raise FORMULA [phi] + in repair_tvar_sorts (do_formula true phi Vartab.empty) end + +fun check_formula ctxt = + Type_Infer.constrain HOLogic.boolT + #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) + + +(**** Translation of TSTP files to Isar Proofs ****) + +fun unvarify_term (Var ((s, 0), T)) = Free (s, T) + | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t]) + +fun decode_line full_types tfrees (Definition (num, phi1, phi2)) ctxt = + let + val thy = ProofContext.theory_of ctxt + val t1 = prop_from_formula thy full_types tfrees phi1 + val vars = snd (strip_comb t1) + val frees = map unvarify_term vars + val unvarify_args = subst_atomic (vars ~~ frees) + val t2 = prop_from_formula thy full_types tfrees phi2 + val (t1, t2) = + HOLogic.eq_const HOLogic.typeT $ t1 $ t2 + |> unvarify_args |> uncombine_term |> check_formula ctxt + |> HOLogic.dest_eq + in + (Definition (num, t1, t2), + fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt) + end + | decode_line full_types tfrees (Inference (num, u, deps)) ctxt = + let + val thy = ProofContext.theory_of ctxt + val t = u |> prop_from_formula thy full_types tfrees + |> uncombine_term |> check_formula ctxt + in + (Inference (num, t, deps), + fold Variable.declare_term (OldTerm.term_frees t) ctxt) + end +fun decode_lines ctxt full_types tfrees lines = + fst (fold_map (decode_line full_types tfrees) lines ctxt) + +fun is_same_inference _ (Definition _) = false + | is_same_inference t (Inference (_, t', _)) = t aconv t' + +(* No "real" literals means only type information (tfree_tcs, clsrel, or + clsarity). *) +val is_only_type_information = curry (op aconv) HOLogic.true_const + +fun replace_one_dep (old, new) dep = if dep = old then new else [dep] +fun replace_deps_in_line _ (line as Definition _) = line + | replace_deps_in_line p (Inference (num, t, deps)) = + Inference (num, t, fold (union (op =) o replace_one_dep p) deps []) + +(* Discard axioms; consolidate adjacent lines that prove the same formula, since + they differ only in type information.*) +fun add_line _ _ (line as Definition _) lines = line :: lines + | add_line conjecture_shape axiom_names (Inference (num, t, [])) lines = + (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or + definitions. *) + if is_axiom_number axiom_names num then + (* Axioms are not proof lines. *) + if is_only_type_information t then + map (replace_deps_in_line (num, [])) lines + (* Is there a repetition? If so, replace later line by earlier one. *) + else case take_prefix (not o is_same_inference t) lines of + (_, []) => lines (*no repetition of proof line*) + | (pre, Inference (num', _, _) :: post) => + pre @ map (replace_deps_in_line (num', [num])) post + else if is_conjecture_number conjecture_shape num then + Inference (num, negate_term t, []) :: lines + else + map (replace_deps_in_line (num, [])) lines + | add_line _ _ (Inference (num, t, deps)) lines = + (* Type information will be deleted later; skip repetition test. *) + if is_only_type_information t then + Inference (num, t, deps) :: lines + (* Is there a repetition? If so, replace later line by earlier one. *) + else case take_prefix (not o is_same_inference t) lines of + (* FIXME: Doesn't this code risk conflating proofs involving different + types? *) + (_, []) => Inference (num, t, deps) :: lines + | (pre, Inference (num', t', _) :: post) => + Inference (num, t', deps) :: + pre @ map (replace_deps_in_line (num', [num])) post + +(* Recursively delete empty lines (type information) from the proof. *) +fun add_nontrivial_line (Inference (num, t, [])) lines = + if is_only_type_information t then delete_dep num lines + else Inference (num, t, []) :: lines + | add_nontrivial_line line lines = line :: lines +and delete_dep num lines = + fold_rev add_nontrivial_line (map (replace_deps_in_line (num, [])) lines) [] + +(* ATPs sometimes reuse free variable names in the strangest ways. Removing + offending lines often does the trick. *) +fun is_bad_free frees (Free x) = not (member (op =) frees x) + | is_bad_free _ _ = false + +(* Vampire is keen on producing these. *) +fun is_trivial_formula (@{const Not} $ (Const (@{const_name HOL.eq}, _) + $ t1 $ t2)) = (t1 aconv t2) + | is_trivial_formula _ = false + +fun add_desired_line _ _ _ _ (line as Definition (num, _, _)) (j, lines) = + (j, line :: map (replace_deps_in_line (num, [])) lines) + | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees + (Inference (num, t, deps)) (j, lines) = + (j + 1, + if is_axiom_number axiom_names num orelse + is_conjecture_number conjecture_shape num orelse + (not (is_only_type_information t) andalso + null (Term.add_tvars t []) andalso + not (exists_subterm (is_bad_free frees) t) andalso + not (is_trivial_formula t) andalso + (null lines orelse (* last line must be kept *) + (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then + Inference (num, t, deps) :: lines (* keep line *) + else + map (replace_deps_in_line (num, deps)) lines) (* drop line *) + +(** EXTRACTING LEMMAS **) + +(* Like "split_line", but ignores "\n" that follow a comma (as in SNARK's + output). *) +val split_proof_lines = + let + fun aux [] [] = [] + | aux line [] = [implode (rev line)] + | aux line ("," :: "\n" :: rest) = aux ("," :: line) rest + | aux line ("\n" :: rest) = aux line [] @ aux [] rest + | aux line (s :: rest) = aux (s :: line) rest + in aux [] o explode end + +(* A list consisting of the first number in each line is returned. For TSTP, + interesting lines have the form "fof(108, axiom, ...)", where the number + (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where + the first number (108) is extracted. For Vampire, we look for + "108. ... [input]". *) +fun used_facts_in_atp_proof axiom_names atp_proof = + let + fun axiom_names_at_index num = + let val j = Int.fromString num |> the_default ~1 in + if is_axiom_number axiom_names j then Vector.sub (axiom_names, j - 1) + else [] + end + val tokens_of = + String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_") + fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) = + if tag = "cnf" orelse tag = "fof" then + (case strip_prefix_and_unascii axiom_prefix (List.last rest) of + SOME name => + if member (op =) rest "file" then + ([(name, name |> find_first_in_list_vector axiom_names |> the)] + handle Option.Option => + error ("No such fact: " ^ quote name ^ ".")) + else + axiom_names_at_index num + | NONE => axiom_names_at_index num) + else + [] + | do_line (num :: "0" :: "Inp" :: _) = axiom_names_at_index num + | do_line (num :: rest) = + (case List.last rest of "input" => axiom_names_at_index num | _ => []) + | do_line _ = [] + in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end + +val indent_size = 2 +val no_label = ("", ~1) + +val raw_prefix = "X" +val assum_prefix = "A" +val fact_prefix = "F" + +fun string_for_label (s, num) = s ^ string_of_int num + +fun metis_using [] = "" + | metis_using ls = + "using " ^ space_implode " " (map string_for_label ls) ^ " " +fun metis_apply _ 1 = "by " + | metis_apply 1 _ = "apply " + | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply " +fun metis_name full_types = if full_types then "metisFT" else "metis" +fun metis_call full_types [] = metis_name full_types + | metis_call full_types ss = + "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")" +fun metis_command full_types i n (ls, ss) = + metis_using ls ^ metis_apply i n ^ metis_call full_types ss +fun metis_line full_types i n ss = + "Try this command: " ^ + Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "." +fun minimize_line _ [] = "" + | minimize_line minimize_command ss = + case minimize_command ss of + "" => "" + | command => + "\nTo minimize the number of lemmas, try this: " ^ + Markup.markup Markup.sendback command ^ "." + +fun used_facts axiom_names = + used_facts_in_atp_proof axiom_names + #> List.partition (curry (op =) Chained o snd) + #> pairself (sort_distinct (string_ord o pairself fst)) + +fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names, + goal, i) = + let + val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof + val n = Logic.count_prems (prop_of goal) + in + (metis_line full_types i n (map fst other_lemmas) ^ + minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)), + other_lemmas @ chained_lemmas) + end + +(** Isar proof construction and manipulation **) + +fun merge_fact_sets (ls1, ss1) (ls2, ss2) = + (union (op =) ls1 ls2, union (op =) ss1 ss2) + +type label = string * int +type facts = label list * string list + +datatype qualifier = Show | Then | Moreover | Ultimately + +datatype step = + Fix of (string * typ) list | + Let of term * term | + Assume of label * term | + Have of qualifier list * label * term * byline +and byline = + ByMetis of facts | + CaseSplit of step list list * facts + +fun smart_case_split [] facts = ByMetis facts + | smart_case_split proofs facts = CaseSplit (proofs, facts) + +fun add_fact_from_dep axiom_names num = + if is_axiom_number axiom_names num then + apsnd (union (op =) (map fst (Vector.sub (axiom_names, num - 1)))) + else + apfst (insert (op =) (raw_prefix, num)) + +fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t +fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t + +fun step_for_line _ _ (Definition (_, t1, t2)) = Let (t1, t2) + | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t) + | step_for_line axiom_names j (Inference (num, t, deps)) = + Have (if j = 1 then [Show] else [], (raw_prefix, num), + forall_vars t, + ByMetis (fold (add_fact_from_dep axiom_names) deps ([], []))) + +fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor + atp_proof conjecture_shape axiom_names params frees = + let + val lines = + atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *) + |> parse_proof pool + |> sort (int_ord o pairself raw_step_number) + |> decode_lines ctxt full_types tfrees + |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names) + |> rpair [] |-> fold_rev add_nontrivial_line + |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor + conjecture_shape axiom_names frees) + |> snd + in + (if null params then [] else [Fix params]) @ + map2 (step_for_line axiom_names) (length lines downto 1) lines + end + +(* When redirecting proofs, we keep information about the labels seen so far in + the "backpatches" data structure. The first component indicates which facts + should be associated with forthcoming proof steps. The second component is a + pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should + become assumptions and "drop_ls" are the labels that should be dropped in a + case split. *) +type backpatches = (label * facts) list * (label list * label list) + +fun used_labels_of_step (Have (_, _, _, by)) = + (case by of + ByMetis (ls, _) => ls + | CaseSplit (proofs, (ls, _)) => + fold (union (op =) o used_labels_of) proofs ls) + | used_labels_of_step _ = [] +and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof [] + +fun new_labels_of_step (Fix _) = [] + | new_labels_of_step (Let _) = [] + | new_labels_of_step (Assume (l, _)) = [l] + | new_labels_of_step (Have (_, l, _, _)) = [l] +val new_labels_of = maps new_labels_of_step + +val join_proofs = + let + fun aux _ [] = NONE + | aux proof_tail (proofs as (proof1 :: _)) = + if exists null proofs then + NONE + else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then + aux (hd proof1 :: proof_tail) (map tl proofs) + else case hd proof1 of + Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *) + if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t') + | _ => false) (tl proofs) andalso + not (exists (member (op =) (maps new_labels_of proofs)) + (used_labels_of proof_tail)) then + SOME (l, t, map rev proofs, proof_tail) + else + NONE + | _ => NONE + in aux [] o map rev end + +fun case_split_qualifiers proofs = + case length proofs of + 0 => [] + | 1 => [Then] + | _ => [Ultimately] + +fun redirect_proof conjecture_shape hyp_ts concl_t proof = + let + (* The first pass outputs those steps that are independent of the negated + conjecture. The second pass flips the proof by contradiction to obtain a + direct proof, introducing case splits when an inference depends on + several facts that depend on the negated conjecture. *) + fun find_hyp num = + nth hyp_ts (index_in_shape num conjecture_shape) + handle Subscript => + raise Fail ("Cannot find hypothesis " ^ Int.toString num) + val concl_ls = map (pair raw_prefix) (List.last conjecture_shape) + val canonicalize_labels = + map (fn l => if member (op =) concl_ls l then hd concl_ls else l) + #> distinct (op =) + fun first_pass ([], contra) = ([], contra) + | first_pass ((step as Fix _) :: proof, contra) = + first_pass (proof, contra) |>> cons step + | first_pass ((step as Let _) :: proof, contra) = + first_pass (proof, contra) |>> cons step + | first_pass ((step as Assume (l as (_, num), _)) :: proof, contra) = + if member (op =) concl_ls l then + first_pass (proof, contra ||> l = hd concl_ls ? cons step) + else + first_pass (proof, contra) |>> cons (Assume (l, find_hyp num)) + | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) = + let + val ls = canonicalize_labels ls + val step = Have (qs, l, t, ByMetis (ls, ss)) + in + if exists (member (op =) (fst contra)) ls then + first_pass (proof, contra |>> cons l ||> cons step) + else + first_pass (proof, contra) |>> cons step + end + | first_pass _ = raise Fail "malformed proof" + val (proof_top, (contra_ls, contra_proof)) = + first_pass (proof, (concl_ls, [])) + val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst + fun backpatch_labels patches ls = + fold merge_fact_sets (map (backpatch_label patches) ls) ([], []) + fun second_pass end_qs ([], assums, patches) = + ([Have (end_qs, no_label, concl_t, + ByMetis (backpatch_labels patches (map snd assums)))], patches) + | second_pass end_qs (Assume (l, t) :: proof, assums, patches) = + second_pass end_qs (proof, (t, l) :: assums, patches) + | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums, + patches) = + if member (op =) (snd (snd patches)) l andalso + not (member (op =) (fst (snd patches)) l) andalso + not (AList.defined (op =) (fst patches) l) then + second_pass end_qs (proof, assums, patches ||> apsnd (append ls)) + else + (case List.partition (member (op =) contra_ls) ls of + ([contra_l], co_ls) => + if member (op =) qs Show then + second_pass end_qs (proof, assums, + patches |>> cons (contra_l, (co_ls, ss))) + else + second_pass end_qs + (proof, assums, + patches |>> cons (contra_l, (l :: co_ls, ss))) + |>> cons (if member (op =) (fst (snd patches)) l then + Assume (l, negate_term t) + else + Have (qs, l, negate_term t, + ByMetis (backpatch_label patches l))) + | (contra_ls as _ :: _, co_ls) => + let + val proofs = + map_filter + (fn l => + if member (op =) concl_ls l then + NONE + else + let + val drop_ls = filter (curry (op <>) l) contra_ls + in + second_pass [] + (proof, assums, + patches ||> apfst (insert (op =) l) + ||> apsnd (union (op =) drop_ls)) + |> fst |> SOME + end) contra_ls + val (assumes, facts) = + if member (op =) (fst (snd patches)) l then + ([Assume (l, negate_term t)], (l :: co_ls, ss)) + else + ([], (co_ls, ss)) + in + (case join_proofs proofs of + SOME (l, t, proofs, proof_tail) => + Have (case_split_qualifiers proofs @ + (if null proof_tail then end_qs else []), l, t, + smart_case_split proofs facts) :: proof_tail + | NONE => + [Have (case_split_qualifiers proofs @ end_qs, no_label, + concl_t, smart_case_split proofs facts)], + patches) + |>> append assumes + end + | _ => raise Fail "malformed proof") + | second_pass _ _ = raise Fail "malformed proof" + val proof_bottom = + second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst + in proof_top @ proof_bottom end + +(* FIXME: Still needed? Probably not. *) +val kill_duplicate_assumptions_in_proof = + let + fun relabel_facts subst = + apfst (map (fn l => AList.lookup (op =) subst l |> the_default l)) + fun do_step (step as Assume (l, t)) (proof, subst, assums) = + (case AList.lookup (op aconv) assums t of + SOME l' => (proof, (l, l') :: subst, assums) + | NONE => (step :: proof, subst, (t, l) :: assums)) + | do_step (Have (qs, l, t, by)) (proof, subst, assums) = + (Have (qs, l, t, + case by of + ByMetis facts => ByMetis (relabel_facts subst facts) + | CaseSplit (proofs, facts) => + CaseSplit (map do_proof proofs, relabel_facts subst facts)) :: + proof, subst, assums) + | do_step step (proof, subst, assums) = (step :: proof, subst, assums) + and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev + in do_proof end + +val then_chain_proof = + let + fun aux _ [] = [] + | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof + | aux l' (Have (qs, l, t, by) :: proof) = + (case by of + ByMetis (ls, ss) => + Have (if member (op =) ls l' then + (Then :: qs, l, t, + ByMetis (filter_out (curry (op =) l') ls, ss)) + else + (qs, l, t, ByMetis (ls, ss))) + | CaseSplit (proofs, facts) => + Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) :: + aux l proof + | aux _ (step :: proof) = step :: aux no_label proof + in aux no_label end + +fun kill_useless_labels_in_proof proof = + let + val used_ls = used_labels_of proof + fun do_label l = if member (op =) used_ls l then l else no_label + fun do_step (Assume (l, t)) = Assume (do_label l, t) + | do_step (Have (qs, l, t, by)) = + Have (qs, do_label l, t, + case by of + CaseSplit (proofs, facts) => + CaseSplit (map (map do_step) proofs, facts) + | _ => by) + | do_step step = step + in map do_step proof end + +fun prefix_for_depth n = replicate_string (n + 1) + +val relabel_proof = + let + fun aux _ _ _ [] = [] + | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) = + if l = no_label then + Assume (l, t) :: aux subst depth (next_assum, next_fact) proof + else + let val l' = (prefix_for_depth depth assum_prefix, next_assum) in + Assume (l', t) :: + aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof + end + | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) = + let + val (l', subst, next_fact) = + if l = no_label then + (l, subst, next_fact) + else + let + val l' = (prefix_for_depth depth fact_prefix, next_fact) + in (l', (l, l') :: subst, next_fact + 1) end + val relabel_facts = + apfst (map (fn l => + case AList.lookup (op =) subst l of + SOME l' => l' + | NONE => raise Fail ("unknown label " ^ + quote (string_for_label l)))) + val by = + case by of + ByMetis facts => ByMetis (relabel_facts facts) + | CaseSplit (proofs, facts) => + CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs, + relabel_facts facts) + in + Have (qs, l', t, by) :: + aux subst depth (next_assum, next_fact) proof + end + | aux subst depth nextp (step :: proof) = + step :: aux subst depth nextp proof + in aux [] 0 (1, 1) end + +fun string_for_proof ctxt full_types i n = + let + fun fix_print_mode f x = + setmp_CRITICAL show_no_free_types true + (setmp_CRITICAL show_types true + (Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) + (print_mode_value ())) f)) x + fun do_indent ind = replicate_string (ind * indent_size) " " + fun do_free (s, T) = + maybe_quote s ^ " :: " ^ + maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T) + fun do_label l = if l = no_label then "" else string_for_label l ^ ": " + fun do_have qs = + (if member (op =) qs Moreover then "moreover " else "") ^ + (if member (op =) qs Ultimately then "ultimately " else "") ^ + (if member (op =) qs Then then + if member (op =) qs Show then "thus" else "hence" + else + if member (op =) qs Show then "show" else "have") + val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt) + fun do_facts (ls, ss) = + metis_command full_types 1 1 + (ls |> sort_distinct (prod_ord string_ord int_ord), + ss |> sort_distinct string_ord) + and do_step ind (Fix xs) = + do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n" + | do_step ind (Let (t1, t2)) = + do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n" + | do_step ind (Assume (l, t)) = + do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n" + | do_step ind (Have (qs, l, t, ByMetis facts)) = + do_indent ind ^ do_have qs ^ " " ^ + do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n" + | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) = + space_implode (do_indent ind ^ "moreover\n") + (map (do_block ind) proofs) ^ + do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^ + do_facts facts ^ "\n" + and do_steps prefix suffix ind steps = + let val s = implode (map (do_step ind) steps) in + replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ + String.extract (s, ind * indent_size, + SOME (size s - ind * indent_size - 1)) ^ + suffix ^ "\n" + end + and do_block ind proof = do_steps "{ " " }" (ind + 1) proof + (* One-step proofs are pointless; better use the Metis one-liner + directly. *) + and do_proof [Have (_, _, _, ByMetis _)] = "" + | do_proof proof = + (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ + do_indent 0 ^ "proof -\n" ^ + do_steps "" "" 1 proof ^ + do_indent 0 ^ (if n <> 1 then "next" else "qed") + in do_proof end + +fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape) + (other_params as (full_types, _, atp_proof, axiom_names, + goal, i)) = + let + val (params, hyp_ts, concl_t) = strip_subgoal goal i + val frees = fold Term.add_frees (concl_t :: hyp_ts) [] + val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) [] + val n = Logic.count_prems (prop_of goal) + val (one_line_proof, lemma_names) = metis_proof_text other_params + fun isar_proof_for () = + case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor + atp_proof conjecture_shape axiom_names params + frees + |> redirect_proof conjecture_shape hyp_ts concl_t + |> kill_duplicate_assumptions_in_proof + |> then_chain_proof + |> kill_useless_labels_in_proof + |> relabel_proof + |> string_for_proof ctxt full_types i n of + "" => "\nNo structured proof available." + | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof + val isar_proof = + if debug then + isar_proof_for () + else + try isar_proof_for () + |> the_default "\nWarning: The Isar proof construction failed." + in (one_line_proof ^ isar_proof, lemma_names) end + +fun proof_text isar_proof isar_params other_params = + (if isar_proof then isar_proof_text isar_params else metis_proof_text) + other_params + +end; diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Thu Sep 02 08:40:25 2010 +0200 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Thu Sep 02 08:51:16 2010 +0200 @@ -9,6 +9,7 @@ signature SLEDGEHAMMER_TRANSLATE = sig type 'a problem = 'a ATP_Problem.problem + type fol_formula val axiom_prefix : string val conjecture_prefix : string @@ -16,9 +17,12 @@ val class_rel_clause_prefix : string val arity_clause_prefix : string val tfrees_name : string + val prepare_axiom : + Proof.context -> (string * 'a) * thm + -> term * ((string * 'a) * fol_formula) option val prepare_problem : Proof.context -> bool -> bool -> bool -> bool -> term list -> term - -> ((string * 'a) * thm) list + -> (term * ((string * 'a) * fol_formula) option) list -> string problem * string Symtab.table * int * (string * 'a) list vector end; @@ -246,20 +250,15 @@ |> map_filter (Option.map snd o make_axiom ctxt false) end -fun prepare_formulas ctxt full_types hyp_ts concl_t axiom_pairs = +fun prepare_axiom ctxt (ax as (_, th)) = (prop_of th, make_axiom ctxt true ax) + +fun prepare_formulas ctxt full_types hyp_ts concl_t axioms = let val thy = ProofContext.theory_of ctxt - val axiom_ts = map (prop_of o snd) axiom_pairs - val hyp_ts = - if null hyp_ts then - [] - else - (* Remove existing axioms from the conjecture, as this can dramatically - boost an ATP's performance (for some reason). *) - let - val axiom_table = fold (Termtab.update o rpair ()) axiom_ts - Termtab.empty - in hyp_ts |> filter_out (Termtab.defined axiom_table) end + val (axiom_ts, prepared_axioms) = ListPair.unzip axioms + (* Remove existing axioms from the conjecture, as this can dramatically + boost an ATP's performance (for some reason). *) + val hyp_ts = hyp_ts |> filter_out (member (op aconv) axiom_ts) val goal_t = Logic.list_implies (hyp_ts, concl_t) val is_FO = Meson.is_fol_term thy goal_t val subs = tfree_classes_of_terms [goal_t] @@ -267,8 +266,7 @@ val tycons = type_consts_of_terms thy (goal_t :: axiom_ts) (* TFrees in the conjecture; TVars in the axioms *) val conjectures = make_conjecture ctxt (hyp_ts @ [concl_t]) - val (axiom_names, axioms) = - ListPair.unzip (map_filter (make_axiom ctxt true) axiom_pairs) + val (axiom_names, axioms) = ListPair.unzip (map_filter I prepared_axioms) val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms val (supers', arity_clauses) = make_arity_clauses thy tycons supers val class_rel_clauses = make_class_rel_clauses thy subs supers' @@ -501,12 +499,12 @@ (const_table_for_problem explicit_apply problem) problem fun prepare_problem ctxt readable_names explicit_forall full_types - explicit_apply hyp_ts concl_t axiom_pairs = + explicit_apply hyp_ts concl_t axioms = let val thy = ProofContext.theory_of ctxt val (axiom_names, (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses)) = - prepare_formulas ctxt full_types hyp_ts concl_t axiom_pairs + prepare_formulas ctxt full_types hyp_ts concl_t axioms val axiom_lines = map (problem_line_for_fact axiom_prefix full_types) axioms val helper_lines = map (problem_line_for_fact helper_prefix full_types) helper_facts diff -r 7cf8beb31e0f -r e820beaf7d8c src/HOL/Tools/async_manager.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/async_manager.ML Thu Sep 02 08:51:16 2010 +0200 @@ -0,0 +1,234 @@ +(* Title: HOL/Tools/async_manager.ML + Author: Fabian Immler, TU Muenchen + Author: Makarius + Author: Jasmin Blanchette, TU Muenchen + +Central manager for asynchronous diagnosis tool threads. +*) + +signature ASYNC_MANAGER = +sig + val launch : + string -> Time.time -> Time.time -> string -> (unit -> string) -> unit + val kill_threads : string -> string -> unit + val running_threads : string -> string -> unit + val thread_messages : string -> string -> int option -> unit +end; + +structure Async_Manager : ASYNC_MANAGER = +struct + +(** preferences **) + +val message_store_limit = 20; +val message_display_limit = 5; + + +(** thread management **) + +(* data structures over threads *) + +structure Thread_Heap = Heap +( + type elem = Time.time * Thread.thread; + fun ord ((a, _), (b, _)) = Time.compare (a, b); +); + +fun lookup_thread xs = AList.lookup Thread.equal xs; +fun delete_thread xs = AList.delete Thread.equal xs; +fun update_thread xs = AList.update Thread.equal xs; + + +(* state of thread manager *) + +type state = + {manager: Thread.thread option, + timeout_heap: Thread_Heap.T, + active: (Thread.thread * (string * Time.time * Time.time * string)) list, + canceling: (Thread.thread * (string * Time.time * string)) list, + messages: (string * string) list, + store: (string * string) list} + +fun make_state manager timeout_heap active canceling messages store : state = + {manager = manager, timeout_heap = timeout_heap, active = active, + canceling = canceling, messages = messages, store = store} + +val global_state = Synchronized.var "async_manager" + (make_state NONE Thread_Heap.empty [] [] [] []); + + +(* unregister thread *) + +fun unregister message thread = + Synchronized.change global_state + (fn state as {manager, timeout_heap, active, canceling, messages, store} => + (case lookup_thread active thread of + SOME (tool, birth_time, _, desc) => + let + val active' = delete_thread thread active; + val now = Time.now () + val canceling' = (thread, (tool, now, desc)) :: canceling + val message' = desc ^ "\n" ^ message + val messages' = (tool, message') :: messages; + val store' = (tool, message') :: + (if length store <= message_store_limit then store + else #1 (chop message_store_limit store)); + in make_state manager timeout_heap active' canceling' messages' store' end + | NONE => state)); + + +(* main manager thread -- only one may exist *) + +val min_wait_time = Time.fromMilliseconds 300; +val max_wait_time = Time.fromSeconds 10; + +fun replace_all bef aft = + let + fun aux seen "" = String.implode (rev seen) + | aux seen s = + if String.isPrefix bef s then + aux seen "" ^ aft ^ aux [] (unprefix bef s) + else + aux (String.sub (s, 0) :: seen) (String.extract (s, 1, NONE)) + in aux [] end + +(* This is a workaround for Proof General's off-by-a-few sendback display bug, + whereby "pr" in "proof" is not highlighted. *) +fun break_into_chunks xs = + maps (space_explode "\000" o replace_all "\n\n" "\000" o snd) xs + +fun print_new_messages () = + case Synchronized.change_result global_state + (fn {manager, timeout_heap, active, canceling, messages, store} => + (messages, make_state manager timeout_heap active canceling [] + store)) of + [] => () + | msgs as (tool, _) :: _ => + let val ss = break_into_chunks msgs in + List.app priority (tool ^ ": " ^ hd ss :: tl ss) + end + +fun check_thread_manager () = Synchronized.change global_state + (fn state as {manager, timeout_heap, active, canceling, messages, store} => + if (case manager of SOME thread => Thread.isActive thread | NONE => false) then state + else let val manager = SOME (Toplevel.thread false (fn () => + let + fun time_limit timeout_heap = + (case try Thread_Heap.min timeout_heap of + NONE => Time.+ (Time.now (), max_wait_time) + | SOME (time, _) => time); + + (*action: find threads whose timeout is reached, and interrupt canceling threads*) + fun action {manager, timeout_heap, active, canceling, messages, store} = + let val (timeout_threads, timeout_heap') = + Thread_Heap.upto (Time.now (), Thread.self ()) timeout_heap; + in + if null timeout_threads andalso null canceling then + NONE + else + let + val _ = List.app (Simple_Thread.interrupt o #1) canceling + val canceling' = filter (Thread.isActive o #1) canceling + val state' = make_state manager timeout_heap' active canceling' messages store; + in SOME (map #2 timeout_threads, state') end + end; + in + while Synchronized.change_result global_state + (fn state as {timeout_heap, active, canceling, messages, store, ...} => + if null active andalso null canceling andalso null messages + then (false, make_state NONE timeout_heap active canceling messages store) + else (true, state)) + do + (Synchronized.timed_access global_state (SOME o time_limit o #timeout_heap) action + |> these + |> List.app (unregister "Timed out.\n"); + print_new_messages (); + (*give threads some time to respond to interrupt*) + OS.Process.sleep min_wait_time) + end)) + in make_state manager timeout_heap active canceling messages store end) + + +(* register thread *) + +fun register tool birth_time death_time desc thread = + (Synchronized.change global_state + (fn {manager, timeout_heap, active, canceling, messages, store} => + let + val timeout_heap' = Thread_Heap.insert (death_time, thread) timeout_heap; + val active' = update_thread (thread, (tool, birth_time, death_time, desc)) active; + val state' = make_state manager timeout_heap' active' canceling messages store; + in state' end); + check_thread_manager ()) + + +fun launch tool birth_time death_time desc f = + (Toplevel.thread true + (fn () => + let + val self = Thread.self () + val _ = register tool birth_time death_time desc self + val message = f () + in unregister message self end); + ()) + + +(** user commands **) + +(* kill threads *) + +fun kill_threads tool workers = Synchronized.change global_state + (fn {manager, timeout_heap, active, canceling, messages, store} => + let + val killing = + map_filter (fn (th, (tool', _, _, desc)) => + if tool' = tool then SOME (th, (tool', Time.now (), desc)) + else NONE) active + val state' = make_state manager timeout_heap [] (killing @ canceling) messages store; + val _ = if null killing then () else priority ("Killed active " ^ workers ^ ".") + in state' end); + + +(* running threads *) + +fun seconds time = string_of_int (Time.toSeconds time) ^ " s" + +fun running_threads tool workers = + let + val {active, canceling, ...} = Synchronized.value global_state; + + val now = Time.now (); + fun running_info (_, (tool', birth_time, death_time, desc)) = + if tool' = tool then + SOME ("Running: " ^ seconds (Time.- (now, birth_time)) ^ " -- " ^ + seconds (Time.- (death_time, now)) ^ " to live:\n" ^ desc) + else + NONE + fun canceling_info (_, (tool', death_time, desc)) = + if tool' = tool then + SOME ("Trying to interrupt thread since " ^ + seconds (Time.- (now, death_time)) ^ ":\n" ^ desc) + else + NONE + val running = + case map_filter running_info active of + [] => ["No " ^ workers ^ " running."] + | ss => "Running " ^ workers ^ ":" :: ss + val interrupting = + case map_filter canceling_info canceling of + [] => [] + | ss => "Trying to interrupt the following " ^ workers ^ ":" :: ss + in priority (space_implode "\n\n" (running @ interrupting)) end + +fun thread_messages tool worker opt_limit = + let + val limit = the_default message_display_limit opt_limit; + val tool_store = Synchronized.value global_state + |> #store |> filter (curry (op =) tool o fst) + val header = + "Recent " ^ worker ^ " messages" ^ + (if length tool_store <= limit then ":" + else " (" ^ string_of_int limit ^ " displayed):"); + in List.app priority (header :: break_into_chunks (#1 (chop limit tool_store))) end + +end;