author | blanchet |
Mon, 06 Jun 2011 22:03:58 +0200 | |
changeset 43217 | 37d507be3014 |
parent 43216 | 7b89836fd607 |
child 43228 | 2ed2f092e990 |
permissions | -rw-r--r-- |
36926 | 1 |
\documentclass[a4paper,12pt]{article} |
2 |
\usepackage[T1]{fontenc} |
|
3 |
\usepackage{amsmath} |
|
4 |
\usepackage{amssymb} |
|
5 |
\usepackage[english,french]{babel} |
|
6 |
\usepackage{color} |
|
7 |
\usepackage{footmisc} |
|
8 |
\usepackage{graphicx} |
|
9 |
%\usepackage{mathpazo} |
|
10 |
\usepackage{multicol} |
|
11 |
\usepackage{stmaryrd} |
|
12 |
%\usepackage[scaled=.85]{beramono} |
|
42511 | 13 |
\usepackage{../../lib/texinputs/isabelle,../iman,../pdfsetup} |
36926 | 14 |
|
43216 | 15 |
\def\qty#1{\ensuremath{\left<\mathit{#1\/}\right>}} |
16 |
\def\qtybf#1{$\mathbf{\left<\textbf{\textit{#1\/}}\right>}$} |
|
17 |
||
36926 | 18 |
%\oddsidemargin=4.6mm |
19 |
%\evensidemargin=4.6mm |
|
20 |
%\textwidth=150mm |
|
21 |
%\topmargin=4.6mm |
|
22 |
%\headheight=0mm |
|
23 |
%\headsep=0mm |
|
24 |
%\textheight=234mm |
|
25 |
||
26 |
\def\Colon{\mathord{:\mkern-1.5mu:}} |
|
27 |
%\def\lbrakk{\mathopen{\lbrack\mkern-3.25mu\lbrack}} |
|
28 |
%\def\rbrakk{\mathclose{\rbrack\mkern-3.255mu\rbrack}} |
|
29 |
\def\lparr{\mathopen{(\mkern-4mu\mid}} |
|
30 |
\def\rparr{\mathclose{\mid\mkern-4mu)}} |
|
31 |
||
32 |
\def\unk{{?}} |
|
33 |
\def\undef{(\lambda x.\; \unk)} |
|
34 |
%\def\unr{\textit{others}} |
|
35 |
\def\unr{\ldots} |
|
36 |
\def\Abs#1{\hbox{\rm{\flqq}}{\,#1\,}\hbox{\rm{\frqq}}} |
|
37 |
\def\Q{{\smash{\lower.2ex\hbox{$\scriptstyle?$}}}} |
|
38 |
||
39 |
\urlstyle{tt} |
|
40 |
||
41 |
\begin{document} |
|
42 |
||
43 |
\selectlanguage{english} |
|
44 |
||
45 |
\title{\includegraphics[scale=0.5]{isabelle_sledgehammer} \\[4ex] |
|
46 |
Hammering Away \\[\smallskipamount] |
|
47 |
\Large A User's Guide to Sledgehammer for Isabelle/HOL} |
|
48 |
\author{\hbox{} \\ |
|
49 |
Jasmin Christian Blanchette \\ |
|
43002
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
50 |
{\normalsize Institut f\"ur Informatik, Technische Universit\"at M\"unchen} \\[4\smallskipamount] |
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
51 |
{\normalsize with contributions from} \\[4\smallskipamount] |
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
52 |
Lawrence C. Paulson \\ |
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
53 |
{\normalsize Computer Laboratory, University of Cambridge} \\ |
36926 | 54 |
\hbox{}} |
55 |
||
56 |
\maketitle |
|
57 |
||
58 |
\tableofcontents |
|
59 |
||
60 |
\setlength{\parskip}{.7em plus .2em minus .1em} |
|
61 |
\setlength{\parindent}{0pt} |
|
62 |
\setlength{\abovedisplayskip}{\parskip} |
|
63 |
\setlength{\abovedisplayshortskip}{.9\parskip} |
|
64 |
\setlength{\belowdisplayskip}{\parskip} |
|
65 |
\setlength{\belowdisplayshortskip}{.9\parskip} |
|
66 |
||
67 |
% General-purpose enum environment with correct spacing |
|
68 |
\newenvironment{enum}% |
|
69 |
{\begin{list}{}{% |
|
70 |
\setlength{\topsep}{.1\parskip}% |
|
71 |
\setlength{\partopsep}{.1\parskip}% |
|
72 |
\setlength{\itemsep}{\parskip}% |
|
73 |
\advance\itemsep by-\parsep}} |
|
74 |
{\end{list}} |
|
75 |
||
76 |
\def\pre{\begingroup\vskip0pt plus1ex\advance\leftskip by\leftmargin |
|
77 |
\advance\rightskip by\leftmargin} |
|
78 |
\def\post{\vskip0pt plus1ex\endgroup} |
|
79 |
||
80 |
\def\prew{\pre\advance\rightskip by-\leftmargin} |
|
81 |
\def\postw{\post} |
|
82 |
||
83 |
\section{Introduction} |
|
84 |
\label{introduction} |
|
85 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
86 |
Sledgehammer is a tool that applies automatic theorem provers (ATPs) |
40942 | 87 |
and satisfiability-modulo-theories (SMT) solvers on the current goal. The |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
88 |
supported ATPs are E \cite{schulz-2002}, LEO-II \cite{leo2}, Satallax |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
89 |
\cite{satallax}, SInE-E \cite{sine}, SNARK \cite{snark}, SPASS |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
90 |
\cite{weidenbach-et-al-2009}, ToFoF-E \cite{tofof}, Vampire |
42945 | 91 |
\cite{riazanov-voronkov-2002}, and Waldmeister \cite{waldmeister}. The ATPs are |
92 |
run either locally or remotely via the System\-On\-TPTP web service |
|
42940 | 93 |
\cite{sutcliffe-2000}. In addition to the ATPs, the SMT solvers Z3 \cite{z3} is |
42945 | 94 |
used by default, and you can tell Sledgehammer to try CVC3 \cite{cvc3} and Yices |
95 |
\cite{yices} as well; these are run either locally or on a server at the TU |
|
96 |
M\"unchen. |
|
36926 | 97 |
|
40073 | 98 |
The problem passed to the automatic provers consists of your current goal |
99 |
together with a heuristic selection of hundreds of facts (theorems) from the |
|
100 |
current theory context, filtered by relevance. Because jobs are run in the |
|
101 |
background, you can continue to work on your proof by other means. Provers can |
|
102 |
be run in parallel. Any reply (which may arrive half a minute later) will appear |
|
103 |
in the Proof General response buffer. |
|
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
104 |
|
40073 | 105 |
The result of a successful proof search is some source text that usually (but |
106 |
not always) reconstructs the proof within Isabelle. For ATPs, the reconstructed |
|
43036 | 107 |
proof relies on the general-purpose Metis prover, which is fully integrated into |
108 |
Isabelle/HOL, with explicit inferences going through the kernel. Thus its |
|
109 |
results are correct by construction. |
|
36926 | 110 |
|
39320 | 111 |
In this manual, we will explicitly invoke the \textbf{sledgehammer} command. |
112 |
Sledgehammer also provides an automatic mode that can be enabled via the |
|
113 |
``Auto Sledgehammer'' option from the ``Isabelle'' menu in Proof General. In |
|
114 |
this mode, Sledgehammer is run on every newly entered theorem. The time limit |
|
115 |
for Auto Sledgehammer and other automatic tools can be set using the ``Auto |
|
116 |
Tools Time Limit'' option. |
|
117 |
||
36926 | 118 |
\newbox\boxA |
119 |
\setbox\boxA=\hbox{\texttt{nospam}} |
|
120 |
||
42763 | 121 |
\newcommand\authoremail{\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@\allowbreak |
122 |
in.\allowbreak tum.\allowbreak de}} |
|
123 |
||
40689 | 124 |
To run Sledgehammer, you must make sure that the theory \textit{Sledgehammer} is |
125 |
imported---this is rarely a problem in practice since it is part of |
|
126 |
\textit{Main}. Examples of Sledgehammer use can be found in Isabelle's |
|
36926 | 127 |
\texttt{src/HOL/Metis\_Examples} directory. |
128 |
Comments and bug reports concerning Sledgehammer or this manual should be |
|
42883 | 129 |
directed to the author at \authoremail. |
36926 | 130 |
|
131 |
\vskip2.5\smallskipamount |
|
132 |
||
133 |
%\textbf{Acknowledgment.} The author would like to thank Mark Summerfield for |
|
134 |
%suggesting several textual improvements. |
|
135 |
||
136 |
\section{Installation} |
|
137 |
\label{installation} |
|
138 |
||
139 |
Sledgehammer is part of Isabelle, so you don't need to install it. However, it |
|
42763 | 140 |
relies on third-party automatic theorem provers (ATPs) and SMT solvers. |
141 |
||
142 |
\subsection{Installing ATPs} |
|
143 |
||
40073 | 144 |
Currently, E, SPASS, and Vampire can be run locally; in addition, E, Vampire, |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
145 |
LEO-II, Satallax, SInE-E, SNARK, ToFoF-E, and Waldmeister are available remotely |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
146 |
via System\-On\-TPTP \cite{sutcliffe-2000}. If you want better performance, you |
42940 | 147 |
should at least install E and SPASS locally. |
36926 | 148 |
|
38043 | 149 |
There are three main ways to install ATPs on your machine: |
36926 | 150 |
|
151 |
\begin{enum} |
|
152 |
\item[$\bullet$] If you installed an official Isabelle package with everything |
|
153 |
inside, it should already include properly setup executables for E and SPASS, |
|
38043 | 154 |
ready to use.% |
155 |
\footnote{Vampire's license prevents us from doing the same for this otherwise |
|
156 |
wonderful tool.} |
|
36926 | 157 |
|
38043 | 158 |
\item[$\bullet$] Alternatively, you can download the Isabelle-aware E and SPASS |
36926 | 159 |
binary packages from Isabelle's download page. Extract the archives, then add a |
41747
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
160 |
line to your \texttt{\$ISABELLE\_HOME\_USER/etc/components}% |
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
161 |
\footnote{The variable \texttt{\$ISABELLE\_HOME\_USER} is set by Isabelle at |
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
162 |
startup. Its value can be retrieved by invoking \texttt{isabelle} |
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
163 |
\texttt{getenv} \texttt{ISABELLE\_HOME\_USER} on the command line.} |
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
164 |
file with the absolute |
40203 | 165 |
path to E or SPASS. For example, if the \texttt{components} does not exist yet |
166 |
and you extracted SPASS to \texttt{/usr/local/spass-3.7}, create the |
|
167 |
\texttt{components} file with the single line |
|
36926 | 168 |
|
169 |
\prew |
|
170 |
\texttt{/usr/local/spass-3.7} |
|
171 |
\postw |
|
172 |
||
38043 | 173 |
in it. |
174 |
||
175 |
\item[$\bullet$] If you prefer to build E or SPASS yourself, or obtained a |
|
176 |
Vampire executable from somewhere (e.g., \url{http://www.vprover.org/}), |
|
177 |
set the environment variable \texttt{E\_HOME}, \texttt{SPASS\_HOME}, or |
|
178 |
\texttt{VAMPIRE\_HOME} to the directory that contains the \texttt{eproof}, |
|
38063 | 179 |
\texttt{SPASS}, or \texttt{vampire} executable. Sledgehammer has been tested |
42845
94c69e441440
mention version 0.6 of Vampire, since that's what's currently available for download
blanchet
parents:
42763
diff
changeset
|
180 |
with E 1.0 and 1.2, SPASS 3.5 and 3.7, and Vampire 0.6 and 1.0% |
38063 | 181 |
\footnote{Following the rewrite of Vampire, the counter for version numbers was |
42845
94c69e441440
mention version 0.6 of Vampire, since that's what's currently available for download
blanchet
parents:
42763
diff
changeset
|
182 |
reset to 0; hence the (new) Vampire versions 0.6 and 1.0 are more recent than, |
94c69e441440
mention version 0.6 of Vampire, since that's what's currently available for download
blanchet
parents:
42763
diff
changeset
|
183 |
say, Vampire 11.5.}% |
38063 | 184 |
. Since the ATPs' output formats are neither documented nor stable, other |
42763 | 185 |
versions of the ATPs might or might not work well with Sledgehammer. Ideally, |
186 |
also set \texttt{E\_VERSION}, \texttt{SPASS\_VERSION}, or |
|
187 |
\texttt{VAMPIRE\_VERSION} to the ATP's version number (e.g., ``1.2''). |
|
36926 | 188 |
\end{enum} |
189 |
||
42763 | 190 |
To check whether E and SPASS are successfully installed, follow the example in |
191 |
\S\ref{first-steps}. If the remote versions of E and SPASS are used (identified |
|
192 |
by the prefix ``\emph{remote\_}''), or if the local versions fail to solve the |
|
193 |
easy goal presented there, this is a sign that something is wrong with your |
|
194 |
installation. |
|
36926 | 195 |
|
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
196 |
Remote ATP invocation via the SystemOnTPTP web service requires Perl with the |
39152
f09b378cb252
make remote ATP invocation work for those people who need to go through a proxy;
blanchet
parents:
38997
diff
changeset
|
197 |
World Wide Web Library (\texttt{libwww-perl}) installed. If you must use a proxy |
f09b378cb252
make remote ATP invocation work for those people who need to go through a proxy;
blanchet
parents:
38997
diff
changeset
|
198 |
server to access the Internet, set the \texttt{http\_proxy} environment variable |
39153 | 199 |
to the proxy, either in the environment in which Isabelle is launched or in your |
41747
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
200 |
\texttt{\char`\~/\$ISABELLE\_HOME\_USER/etc/settings} file. Here are a few examples: |
39152
f09b378cb252
make remote ATP invocation work for those people who need to go through a proxy;
blanchet
parents:
38997
diff
changeset
|
201 |
|
f09b378cb252
make remote ATP invocation work for those people who need to go through a proxy;
blanchet
parents:
38997
diff
changeset
|
202 |
\prew |
39153 | 203 |
\texttt{http\_proxy=http://proxy.example.org} \\ |
204 |
\texttt{http\_proxy=http://proxy.example.org:8080} \\ |
|
205 |
\texttt{http\_proxy=http://joeblow:pAsSwRd@proxy.example.org} |
|
39152
f09b378cb252
make remote ATP invocation work for those people who need to go through a proxy;
blanchet
parents:
38997
diff
changeset
|
206 |
\postw |
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
207 |
|
42763 | 208 |
\subsection{Installing SMT Solvers} |
209 |
||
42945 | 210 |
CVC3, Yices, and Z3 can be run locally or (for CVC3 and Z3) remotely on a TU |
211 |
M\"unchen server. If you want better performance and get the ability to replay |
|
212 |
proofs that rely on the \emph{smt} proof method, you should at least install Z3 |
|
213 |
locally. |
|
42763 | 214 |
|
215 |
There are two main ways of installing SMT solvers locally. |
|
216 |
||
217 |
\begin{enum} |
|
218 |
\item[$\bullet$] If you installed an official Isabelle package with everything |
|
219 |
inside, it should already include properly setup executables for CVC3 and Z3, |
|
220 |
ready to use.% |
|
221 |
\footnote{Yices's license prevents us from doing the same for this otherwise |
|
222 |
wonderful tool.} |
|
223 |
For Z3, you additionally need to set the environment variable |
|
224 |
\texttt{Z3\_NON\_COMMERCIAL} to ``yes'' to confirm that you are a noncommercial |
|
225 |
user. |
|
226 |
||
227 |
\item[$\bullet$] Otherwise, follow the instructions documented in the \emph{SMT} |
|
228 |
theory (\texttt{\$ISABELLE\_HOME/src/HOL/SMT.thy}). |
|
229 |
\end{enum} |
|
230 |
||
36926 | 231 |
\section{First Steps} |
232 |
\label{first-steps} |
|
233 |
||
234 |
To illustrate Sledgehammer in context, let us start a theory file and |
|
235 |
attempt to prove a simple lemma: |
|
236 |
||
237 |
\prew |
|
238 |
\textbf{theory}~\textit{Scratch} \\ |
|
239 |
\textbf{imports}~\textit{Main} \\ |
|
240 |
\textbf{begin} \\[2\smallskipamount] |
|
241 |
% |
|
42945 | 242 |
\textbf{lemma} ``$[a] = [b] \,\Longrightarrow\, a = b$'' \\ |
36926 | 243 |
\textbf{sledgehammer} |
244 |
\postw |
|
245 |
||
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
246 |
Instead of issuing the \textbf{sledgehammer} command, you can also find |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
247 |
Sledgehammer in the ``Commands'' submenu of the ``Isabelle'' menu in Proof |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
248 |
General or press the Emacs key sequence C-c C-a C-s. |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
249 |
Either way, Sledgehammer produces the following output after a few seconds: |
36926 | 250 |
|
251 |
\prew |
|
252 |
\slshape |
|
43035 | 253 |
Sledgehammer: ``\textit{e}'' on goal \\ |
42945 | 254 |
$[a] = [b] \,\Longrightarrow\, a = b$ \\ |
43054 | 255 |
Try this: \textbf{by} (\textit{metis last\_ConsL}) (64 ms). \\[3\smallskipamount] |
42945 | 256 |
% |
43035 | 257 |
Sledgehammer: ``\textit{vampire}'' on goal \\ |
42945 | 258 |
$[a] = [b] \,\Longrightarrow\, a = b$ \\ |
43054 | 259 |
Try this: \textbf{by} (\textit{metis hd.simps}) (14 ms). \\[3\smallskipamount] |
36926 | 260 |
% |
43035 | 261 |
Sledgehammer: ``\textit{spass}'' on goal \\ |
42945 | 262 |
$[a] = [b] \,\Longrightarrow\, a = b$ \\ |
43054 | 263 |
Try this: \textbf{by} (\textit{metis list.inject}) (17 ms). \\[3\smallskipamount] |
36926 | 264 |
% |
43035 | 265 |
Sledgehammer: ``\textit{remote\_waldmeister}'' on goal \\ |
43010
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
266 |
$[a] = [b] \,\Longrightarrow\, a = b$ \\ |
43054 | 267 |
Try this: \textbf{by} (\textit{metis hd.simps}) (15 ms). \\[3\smallskipamount] |
40073 | 268 |
% |
43035 | 269 |
Sledgehammer: ``\textit{remote\_sine\_e}'' on goal \\ |
42945 | 270 |
$[a] = [b] \,\Longrightarrow\, a = b$ \\ |
43054 | 271 |
Try this: \textbf{by} (\textit{metis hd.simps}) (18 ms). \\[3\smallskipamount] |
40942 | 272 |
% |
43035 | 273 |
Sledgehammer: ``\textit{remote\_z3}'' on goal \\ |
42945 | 274 |
$[a] = [b] \,\Longrightarrow\, a = b$ \\ |
43054 | 275 |
Try this: \textbf{by} (\textit{metis list.inject}) (20 ms). |
36926 | 276 |
\postw |
277 |
||
43010
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
278 |
Sledgehammer ran E, SInE-E, SPASS, Vampire, Waldmeister, and Z3 in parallel. |
42945 | 279 |
Depending on which provers are installed and how many processor cores are |
280 |
available, some of the provers might be missing or present with a |
|
43010
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
281 |
\textit{remote\_} prefix. Waldmeister is run only for unit equational problems, |
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
282 |
where the goal's conclusion is a (universally quantified) equation. |
36926 | 283 |
|
43036 | 284 |
For each successful prover, Sledgehammer gives a one-liner proof that uses Metis |
43054 | 285 |
or the \textit{smt} proof method. For Metis, approximate timings are shown in |
286 |
parentheses, indicating how fast the call is. You can click the proof to insert |
|
287 |
it into the theory text. |
|
36926 | 288 |
|
43054 | 289 |
In addition, you can ask Sledgehammer for an Isar text proof by passing the |
42883 | 290 |
\textit{isar\_proof} option (\S\ref{output-format}): |
36926 | 291 |
|
292 |
\prew |
|
293 |
\textbf{sledgehammer} [\textit{isar\_proof}] |
|
294 |
\postw |
|
295 |
||
296 |
When Isar proof construction is successful, it can yield proofs that are more |
|
43036 | 297 |
readable and also faster than the Metis one-liners. This feature is experimental |
298 |
and is only available for ATPs. |
|
36926 | 299 |
|
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
300 |
\section{Hints} |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
301 |
\label{hints} |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
302 |
|
42884 | 303 |
This section presents a few hints that should help you get the most out of |
304 |
Sledgehammer and Metis. Frequently (and infrequently) asked questions are |
|
305 |
answered in \S\ref{frequently-asked-questions}. |
|
306 |
||
42945 | 307 |
\newcommand\point[1]{\medskip\par{\sl\bfseries#1}\par\nopagebreak} |
42763 | 308 |
|
309 |
\point{Presimplify the goal} |
|
310 |
||
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
311 |
For best results, first simplify your problem by calling \textit{auto} or at |
42945 | 312 |
least \textit{safe} followed by \textit{simp\_all}. The SMT solvers provide |
313 |
arithmetic decision procedures, but the ATPs typically do not (or if they do, |
|
314 |
Sledgehammer does not use it yet). Apart from Waldmeister, they are not |
|
315 |
especially good at heavy rewriting, but because they regard equations as |
|
316 |
undirected, they often prove theorems that require the reverse orientation of a |
|
317 |
\textit{simp} rule. Higher-order problems can be tackled, but the success rate |
|
318 |
is better for first-order problems. Hence, you may get better results if you |
|
319 |
first simplify the problem to remove higher-order features. |
|
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
320 |
|
42763 | 321 |
\point{Make sure at least E, SPASS, Vampire, and Z3 are installed} |
322 |
||
323 |
Locally installed provers are faster and more reliable than those running on |
|
324 |
servers. See \S\ref{installation} for details on how to install them. |
|
325 |
||
326 |
\point{Familiarize yourself with the most important options} |
|
327 |
||
328 |
Sledgehammer's options are fully documented in \S\ref{command-syntax}. Many of |
|
329 |
the options are very specialized, but serious users of the tool should at least |
|
330 |
familiarize themselves with the following options: |
|
331 |
||
332 |
\begin{enum} |
|
42884 | 333 |
\item[$\bullet$] \textbf{\textit{provers}} (\S\ref{mode-of-operation}) specifies |
334 |
the automatic provers (ATPs and SMT solvers) that should be run whenever |
|
335 |
Sledgehammer is invoked (e.g., ``\textit{provers}~= \textit{e spass |
|
43014 | 336 |
remote\_vampire}''). For convenience, you can omit ``\textit{provers}~='' |
337 |
and simply write the prover names as a space-separated list (e.g., ``\textit{e |
|
338 |
spass remote\_vampire}''). |
|
42763 | 339 |
|
42884 | 340 |
\item[$\bullet$] \textbf{\textit{full\_types}} (\S\ref{problem-encoding}) |
341 |
specifies whether type-sound encodings should be used. By default, Sledgehammer |
|
342 |
employs a mixture of type-sound and type-unsound encodings, occasionally |
|
43007 | 343 |
yielding unsound ATP proofs. In contrast, SMT solver proofs should always be |
344 |
sound. |
|
42763 | 345 |
|
42884 | 346 |
\item[$\bullet$] \textbf{\textit{max\_relevant}} (\S\ref{relevance-filter}) |
347 |
specifies the maximum number of facts that should be passed to the provers. By |
|
348 |
default, the value is prover-dependent but varies between about 150 and 1000. If |
|
349 |
the provers time out, you can try lowering this value to, say, 100 or 50 and see |
|
350 |
if that helps. |
|
42763 | 351 |
|
42884 | 352 |
\item[$\bullet$] \textbf{\textit{isar\_proof}} (\S\ref{output-format}) specifies |
353 |
that Isar proofs should be generated, instead of one-liner Metis proofs. The |
|
354 |
length of the Isar proofs can be controlled by setting |
|
355 |
\textit{isar\_shrink\_factor} (\S\ref{output-format}). |
|
43038 | 356 |
|
357 |
\item[$\bullet$] \textbf{\textit{timeout}} (\S\ref{timeouts}) controls the |
|
358 |
provers' time limit. It is set to 30 seconds, but since Sledgehammer runs |
|
359 |
asynchronously you should not hesitate to raise this limit to 60 or 120 seconds |
|
360 |
if you are the kind of user who can think clearly while ATPs are active. |
|
42763 | 361 |
\end{enum} |
362 |
||
42884 | 363 |
Options can be set globally using \textbf{sledgehammer\_params} |
43010
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
364 |
(\S\ref{command-syntax}). The command also prints the list of all available |
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
365 |
options with their current value. Fact selection can be influenced by specifying |
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
366 |
``$(\textit{add}{:}~\textit{my\_facts})$'' after the \textbf{sledgehammer} call |
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
367 |
to ensure that certain facts are included, or simply ``$(\textit{my\_facts})$'' |
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
368 |
to force Sledgehammer to run only with $\textit{my\_facts}$. |
42763 | 369 |
|
370 |
\section{Frequently Asked Questions} |
|
371 |
\label{frequently-asked-questions} |
|
372 |
||
42945 | 373 |
This sections answers frequently (and infrequently) asked questions about |
374 |
Sledgehammer. It is a good idea to skim over it now even if you don't have any |
|
375 |
questions at this stage. And if you have any further questions not listed here, |
|
376 |
send them to the author at \authoremail. |
|
377 |
||
42763 | 378 |
\point{Why does Metis fail to reconstruct the proof?} |
379 |
||
42883 | 380 |
There are many reasons. If Metis runs seemingly forever, that is a sign that the |
43036 | 381 |
proof is too difficult for it. Metis's search is complete, so it should |
382 |
eventually find it, but that's little consolation. There are several possible |
|
383 |
solutions: |
|
42763 | 384 |
|
385 |
\begin{enum} |
|
42883 | 386 |
\item[$\bullet$] Try the \textit{isar\_proof} option (\S\ref{output-format}) to |
387 |
obtain a step-by-step Isar proof where each step is justified by Metis. Since |
|
388 |
the steps are fairly small, Metis is more likely to be able to replay them. |
|
42763 | 389 |
|
43036 | 390 |
\item[$\bullet$] Try the \textit{smt} proof method instead of Metis. It is |
391 |
usually stronger, but you need to have Z3 available to replay the proofs, trust |
|
392 |
the SMT solver, or use certificates. See the documentation in the \emph{SMT} |
|
393 |
theory (\texttt{\$ISABELLE\_HOME/src/HOL/SMT.thy}) for details. |
|
42763 | 394 |
|
395 |
\item[$\bullet$] Try the \textit{blast} or \textit{auto} proof methods, passing |
|
43010
a14cf580a5a5
readded Waldmeister as default to the documentation and other minor changes
blanchet
parents:
43008
diff
changeset
|
396 |
the necessary facts via \textbf{unfolding}, \textbf{using}, \textit{intro}{:}, |
42763 | 397 |
\textit{elim}{:}, \textit{dest}{:}, or \textit{simp}{:}, as appropriate. |
398 |
\end{enum} |
|
399 |
||
43036 | 400 |
In some rare cases, Metis fails fairly quickly, and you get the error message |
401 |
||
402 |
\prew |
|
403 |
\slshape |
|
404 |
Proof reconstruction failed. |
|
405 |
\postw |
|
406 |
||
407 |
This usually indicates that Sledgehammer found a type-incorrect proof. |
|
408 |
Sledgehammer erases some type information to speed up the search. Try |
|
409 |
Sledgehammer again with full type information: \textit{full\_types} |
|
410 |
(\S\ref{problem-encoding}), or choose a specific type encoding with |
|
411 |
\textit{type\_sys} (\S\ref{problem-encoding}). Older versions of Sledgehammer |
|
412 |
were frequent victims of this problem. Now this should very seldom be an issue, |
|
413 |
but if you notice many unsound proofs, contact the author at \authoremail. |
|
42883 | 414 |
|
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
415 |
\point{How can I tell whether a generated proof is sound?} |
42883 | 416 |
|
43036 | 417 |
First, if Metis can reconstruct it, the proof is sound (modulo soundness of |
418 |
Isabelle's inference kernel). If it fails or runs seemingly forever, you can try |
|
42883 | 419 |
|
420 |
\prew |
|
421 |
\textbf{apply}~\textbf{--} \\ |
|
422 |
\textbf{sledgehammer} [\textit{type\_sys} = \textit{poly\_tags}] (\textit{metis\_facts}) |
|
423 |
\postw |
|
424 |
||
425 |
where \textit{metis\_facts} is the list of facts appearing in the suggested |
|
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
426 |
Metis call. The automatic provers should be able to re-find the proof very |
42883 | 427 |
quickly if it is sound, and the \textit{type\_sys} $=$ \textit{poly\_tags} |
428 |
option (\S\ref{problem-encoding}) ensures that no unsound proofs are found. |
|
429 |
||
430 |
The \textit{full\_types} option (\S\ref{problem-encoding}) can also be used |
|
431 |
here, but it is unsound in extremely rare degenerate cases such as the |
|
432 |
following: |
|
433 |
||
434 |
\prew |
|
42945 | 435 |
\textbf{lemma} ``$\forall x\> y\Colon{'}\!a.\ x = y \,\Longrightarrow \exists f\> g\Colon\mathit{nat} \Rightarrow {'}\!a.\ f \not= g$'' \\ |
42883 | 436 |
\textbf{sledgehammer} [\textit{full\_types}] (\textit{nat.distinct\/}(1)) |
437 |
\postw |
|
438 |
||
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
439 |
\point{Which facts are passed to the automatic provers?} |
42883 | 440 |
|
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
441 |
The relevance filter assigns a score to every available fact (lemma, theorem, |
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
442 |
definition, or axiom)\ based upon how many constants that fact shares with the |
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
443 |
conjecture. This process iterates to include facts relevant to those just |
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
444 |
accepted, but with a decay factor to ensure termination. The constants are |
42883 | 445 |
weighted to give unusual ones greater significance. The relevance filter copes |
446 |
best when the conjecture contains some unusual constants; if all the constants |
|
447 |
are common, it is unable to discriminate among the hundreds of facts that are |
|
448 |
picked up. The relevance filter is also memoryless: It has no information about |
|
449 |
how many times a particular fact has been used in a proof, and it cannot learn. |
|
42763 | 450 |
|
42883 | 451 |
The number of facts included in a problem varies from prover to prover, since |
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
452 |
some provers get overwhelmed more easily than others. You can show the number of |
42883 | 453 |
facts given using the \textit{verbose} option (\S\ref{output-format}) and the |
454 |
actual facts using \textit{debug} (\S\ref{output-format}). |
|
455 |
||
456 |
Sledgehammer is good at finding short proofs combining a handful of existing |
|
457 |
lemmas. If you are looking for longer proofs, you must typically restrict the |
|
42884 | 458 |
number of facts, by setting the \textit{max\_relevant} option |
459 |
(\S\ref{relevance-filter}) to, say, 50 or 100. |
|
42883 | 460 |
|
42996 | 461 |
You can also influence which facts are actually selected in a number of ways. If |
462 |
you simply want to ensure that a fact is included, you can specify it using the |
|
463 |
``$(\textit{add}{:}~\textit{my\_facts})$'' syntax. For example: |
|
464 |
% |
|
465 |
\prew |
|
466 |
\textbf{sledgehammer} (\textit{add}: \textit{hd.simps} \textit{tl.simps}) |
|
467 |
\postw |
|
468 |
% |
|
469 |
The specified facts then replace the least relevant facts that would otherwise be |
|
470 |
included; the other selected facts remain the same. |
|
471 |
If you want to direct the selection in a particular direction, you can specify |
|
472 |
the facts via \textbf{using}: |
|
473 |
% |
|
474 |
\prew |
|
475 |
\textbf{using} \textit{hd.simps} \textit{tl.simps} \\ |
|
476 |
\textbf{sledgehammer} |
|
477 |
\postw |
|
478 |
% |
|
479 |
The facts are then more likely to be selected than otherwise, and if they are |
|
480 |
selected at iteration $j$ they also influence which facts are selected at |
|
481 |
iterations $j + 1$, $j + 2$, etc. To give them even more weight, try |
|
482 |
% |
|
483 |
\prew |
|
484 |
\textbf{using} \textit{hd.simps} \textit{tl.simps} \\ |
|
485 |
\textbf{apply}~\textbf{--} \\ |
|
486 |
\textbf{sledgehammer} |
|
487 |
\postw |
|
488 |
||
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
489 |
\point{Why are the generated Isar proofs so ugly/detailed/broken?} |
42883 | 490 |
|
491 |
The current implementation is experimental and explodes exponentially in the |
|
492 |
worst case. Work on a new implementation has begun. There is a large body of |
|
493 |
research into transforming resolution proofs into natural deduction proofs (such |
|
494 |
as Isar proofs), which we hope to leverage. In the meantime, a workaround is to |
|
495 |
set the \textit{isar\_shrink\_factor} option (\S\ref{output-format}) to a larger |
|
496 |
value or to try several provers and keep the nicest-looking proof. |
|
497 |
||
43203
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
498 |
\point{What is metisFT?} |
42883 | 499 |
|
500 |
The \textit{metisFT} proof method is the fully-typed version of Metis. It is |
|
501 |
much slower than \textit{metis}, but the proof search is fully typed, and it |
|
502 |
also includes more powerful rules such as the axiom ``$x = \mathit{True} |
|
503 |
\mathrel{\lor} x = \mathit{False}$'' for reasoning in higher-order places (e.g., |
|
43203
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
504 |
in set comprehensions). The method kicks in automatically as a fallback when |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
505 |
\textit{metis} fails, and it is sometimes generated by Sledgehammer instead of |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
506 |
\textit{metis} if the proof obviously requires type information or if |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
507 |
\textit{metis} failed when Sledgehammer preplayed the proof. (By default, |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
508 |
Sledgehammer tries to run \textit{metis} and/or \textit{metisFT} for 4 seconds |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
509 |
to ensure that the generated one-line proofs actually work and to display timing |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
510 |
information. This can be configured using the \textit{preplay\_timeout} option |
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
511 |
(\S\ref{timeouts}).) |
42883 | 512 |
|
513 |
If you see the warning |
|
514 |
||
515 |
\prew |
|
43007 | 516 |
\slshape |
43203
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
517 |
Metis: Falling back on ``\textit{metisFT\/}''. |
42883 | 518 |
\postw |
519 |
||
520 |
in a successful Metis proof, you can advantageously replace the \textit{metis} |
|
43203
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
521 |
call with \textit{metisFT}. |
42850
c8709be8a40f
distinguish between a soft timeout (30 s by defalt) and a hard timeout (60 s), to let minimization-based provers (such as CVC3, Yices, and occasionally the other provers) do their job
blanchet
parents:
42846
diff
changeset
|
522 |
|
43054 | 523 |
\point{Are generated proofs minimal?} |
43036 | 524 |
|
43054 | 525 |
Automatic provers frequently use many more facts than are necessary. |
526 |
Sledgehammer inclues a minimization tool that takes a set of facts returned by a |
|
527 |
given prover and repeatedly calls the same prover or Metis with subsets of those |
|
528 |
axioms in order to find a minimal set. Reducing the number of axioms typically |
|
529 |
improves Metis's speed and success rate, while also removing superfluous clutter |
|
530 |
from the proof scripts. |
|
43036 | 531 |
|
43054 | 532 |
In earlier versions of Sledgehammer, generated proofs were accompanied by a |
533 |
suggestion to invoke the minimization tool. This step is now performed |
|
534 |
implicitly if it can be done in a reasonable amount of time (something that can |
|
535 |
be guessed from the number of facts in the original proof and the time it took |
|
536 |
to find it or replay it). |
|
43036 | 537 |
|
43054 | 538 |
In addition, some provers (notably CVC3, Satallax, and Yices) do not provide |
539 |
proofs or sometimes provide incomplete proofs. The minimizer is then invoked to |
|
540 |
find out which facts are actually needed from the (large) set of facts that was |
|
541 |
initinally given to the prover. Finally, if a prover returns a proof with lots |
|
542 |
of facts, the minimizer is invoked automatically since Metis would be unlikely |
|
543 |
to re-find the proof. |
|
43036 | 544 |
|
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
545 |
\point{A strange error occurred---what should I do?} |
42763 | 546 |
|
547 |
Sledgehammer tries to give informative error messages. Please report any strange |
|
42883 | 548 |
error to the author at \authoremail. This applies double if you get the message |
42763 | 549 |
|
42883 | 550 |
\prew |
42763 | 551 |
\slshape |
42877 | 552 |
The prover found a type-unsound proof involving ``\textit{foo}'', |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43002
diff
changeset
|
553 |
``\textit{bar}'', and ``\textit{baz}'' even though a supposedly type-sound |
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43002
diff
changeset
|
554 |
encoding was used (or, less likely, your axioms are inconsistent). You might |
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43002
diff
changeset
|
555 |
want to report this to the Isabelle developers. |
42883 | 556 |
\postw |
42763 | 557 |
|
558 |
\point{Auto can solve it---why not Sledgehammer?} |
|
559 |
||
560 |
Problems can be easy for \textit{auto} and difficult for automatic provers, but |
|
561 |
the reverse is also true, so don't be discouraged if your first attempts fail. |
|
39320 | 562 |
Because the system refers to all theorems known to Isabelle, it is particularly |
563 |
suitable when your goal has a short proof from lemmas that you don't know about. |
|
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
564 |
|
42883 | 565 |
\point{Why are there so many options?} |
566 |
||
567 |
Sledgehammer's philosophy should work out of the box, without user guidance. |
|
568 |
Many of the options are meant to be used mostly by the Sledgehammer developers |
|
569 |
for experimentation purposes. Of course, feel free to experiment with them if |
|
570 |
you are so inclined. |
|
571 |
||
36926 | 572 |
\section{Command Syntax} |
573 |
\label{command-syntax} |
|
574 |
||
575 |
Sledgehammer can be invoked at any point when there is an open goal by entering |
|
576 |
the \textbf{sledgehammer} command in the theory file. Its general syntax is as |
|
577 |
follows: |
|
578 |
||
579 |
\prew |
|
43216 | 580 |
\textbf{sledgehammer} \qty{subcommand}$^?$ \qty{options}$^?$ \qty{facts\_override}$^?$ \qty{num}$^?$ |
36926 | 581 |
\postw |
582 |
||
583 |
For convenience, Sledgehammer is also available in the ``Commands'' submenu of |
|
584 |
the ``Isabelle'' menu in Proof General or by pressing the Emacs key sequence C-c |
|
585 |
C-a C-s. This is equivalent to entering the \textbf{sledgehammer} command with |
|
586 |
no arguments in the theory text. |
|
587 |
||
43216 | 588 |
In the general syntax, the \qty{subcommand} may be any of the following: |
36926 | 589 |
|
590 |
\begin{enum} |
|
40203 | 591 |
\item[$\bullet$] \textbf{\textit{run} (the default):} Runs Sledgehammer on |
43216 | 592 |
subgoal number \qty{num} (1 by default), with the given options and facts. |
36926 | 593 |
|
43216 | 594 |
\item[$\bullet$] \textbf{\textit{min}:} Attempts to minimize the facts |
595 |
specified in the \qty{facts\_override} argument to obtain a simpler proof |
|
36926 | 596 |
involving fewer facts. The options and goal number are as for \textit{run}. |
597 |
||
40203 | 598 |
\item[$\bullet$] \textbf{\textit{messages}:} Redisplays recent messages issued |
599 |
by Sledgehammer. This allows you to examine results that might have been lost |
|
43216 | 600 |
due to Sledgehammer's asynchronous nature. The \qty{num} argument specifies a |
36926 | 601 |
limit on the number of messages to display (5 by default). |
602 |
||
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41724
diff
changeset
|
603 |
\item[$\bullet$] \textbf{\textit{supported\_provers}:} Prints the list of |
41724 | 604 |
automatic provers supported by Sledgehammer. See \S\ref{installation} and |
605 |
\S\ref{mode-of-operation} for more information on how to install automatic |
|
606 |
provers. |
|
36926 | 607 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
608 |
\item[$\bullet$] \textbf{\textit{running\_provers}:} Prints information about |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
609 |
currently running automatic provers, including elapsed runtime and remaining |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
610 |
time until timeout. |
36926 | 611 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
612 |
\item[$\bullet$] \textbf{\textit{kill\_provers}:} Terminates all running |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
613 |
automatic provers. |
36926 | 614 |
|
615 |
\item[$\bullet$] \textbf{\textit{refresh\_tptp}:} Refreshes the list of remote |
|
616 |
ATPs available at System\-On\-TPTP \cite{sutcliffe-2000}. |
|
617 |
\end{enum} |
|
618 |
||
43216 | 619 |
Sledgehammer's behavior can be influenced by various \qty{options}, which can be |
620 |
specified in brackets after the \textbf{sledgehammer} command. The |
|
621 |
\qty{options} are a list of key--value pairs of the form ``[$k_1 = v_1, |
|
36926 | 622 |
\ldots, k_n = v_n$]''. For Boolean options, ``= \textit{true}'' is optional. For |
623 |
example: |
|
624 |
||
625 |
\prew |
|
43216 | 626 |
\textbf{sledgehammer} [\textit{isar\_proof}, \,\textit{timeout} = 120] |
36926 | 627 |
\postw |
628 |
||
629 |
Default values can be set using \textbf{sledgehammer\_\allowbreak params}: |
|
630 |
||
631 |
\prew |
|
43216 | 632 |
\textbf{sledgehammer\_params} \qty{options} |
36926 | 633 |
\postw |
634 |
||
635 |
The supported options are described in \S\ref{option-reference}. |
|
636 |
||
43216 | 637 |
The \qty{facts\_override} argument lets you alter the set of facts that go |
638 |
through the relevance filter. It may be of the form ``(\qty{facts})'', where |
|
639 |
\qty{facts} is a space-separated list of Isabelle facts (theorems, local |
|
36926 | 640 |
assumptions, etc.), in which case the relevance filter is bypassed and the given |
43216 | 641 |
facts are used. It may also be of the form ``(\textit{add}:\ \qty{facts\/_{\mathrm{1}}})'', |
642 |
``(\textit{del}:\ \qty{facts\/_{\mathrm{2}}})'', or ``(\textit{add}:\ \qty{facts\/_{\mathrm{1}}}\ |
|
643 |
\textit{del}:\ \qty{facts\/_{\mathrm{2}}})'', where the relevance filter is instructed to |
|
644 |
proceed as usual except that it should consider \qty{facts\/_{\mathrm{1}}} |
|
645 |
highly-relevant and \qty{facts\/_{\mathrm{2}}} fully irrelevant. |
|
36926 | 646 |
|
39320 | 647 |
You can instruct Sledgehammer to run automatically on newly entered theorems by |
648 |
enabling the ``Auto Sledgehammer'' option from the ``Isabelle'' menu in Proof |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
649 |
General. For automatic runs, only the first prover set using \textit{provers} |
42736
8005fc9b65ec
ensure that Auto Sledgehammer is run with full type information
blanchet
parents:
42724
diff
changeset
|
650 |
(\S\ref{mode-of-operation}) is considered, fewer facts are passed to the prover, |
43038 | 651 |
\textit{slicing} (\S\ref{mode-of-operation}) is disabled, \textit{full\_types} |
652 |
(\S\ref{problem-encoding}) is enabled, \textit{verbose} (\S\ref{output-format}) |
|
653 |
and \textit{debug} (\S\ref{output-format}) are disabled, and \textit{timeout} |
|
654 |
(\S\ref{timeouts}) is superseded by the ``Auto Tools Time Limit'' in Proof |
|
655 |
General's ``Isabelle'' menu. Sledgehammer's output is also more concise. |
|
39320 | 656 |
|
43216 | 657 |
The \textit{metis} proof method has the syntax |
658 |
||
659 |
\prew |
|
660 |
\textbf{\textit{metis}}~(\qty{type\_sys})${}^?$~\qty{facts}${}^?$ |
|
661 |
\postw |
|
662 |
||
663 |
where \qty{type\_sys} is a type system specification with the same semantics as |
|
664 |
Sledgehammer's \textit{type\_sys} option (\S\ref{problem-encoding}) and |
|
665 |
\qty{facts} is a list of arbitrary facts. \qty{type\_sys} may also be |
|
666 |
\textit{full\_types}, in which case an appropriate type-sound encoding is |
|
667 |
chosen. The companion proof method \textit{metisFT} is an abbreviation for |
|
668 |
``\textit{metis}~(\textit{full\_types})''. |
|
669 |
||
36926 | 670 |
\section{Option Reference} |
671 |
\label{option-reference} |
|
672 |
||
43014 | 673 |
\def\defl{\{} |
674 |
\def\defr{\}} |
|
675 |
||
36926 | 676 |
\def\flushitem#1{\item[]\noindent\kern-\leftmargin \textbf{#1}} |
43014 | 677 |
\def\optrue#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\enskip \defl\textit{true}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
678 |
\def\opfalse#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\enskip \defl\textit{false}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
|
679 |
\def\opsmart#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{smart\_bool}$\bigr]$\enskip \defl\textit{smart}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
|
36926 | 680 |
\def\opnodefault#1#2{\flushitem{\textit{#1} = \qtybf{#2}} \nopagebreak\\[\parskip]} |
43014 | 681 |
\def\opnodefaultbrk#1#2{\flushitem{$\bigl[$\textit{#1} =$\bigr]$ \qtybf{#2}} \nopagebreak\\[\parskip]} |
682 |
\def\opdefault#1#2#3{\flushitem{\textit{#1} = \qtybf{#2}\enskip \defl\textit{#3}\defr} \nopagebreak\\[\parskip]} |
|
36926 | 683 |
\def\oparg#1#2#3{\flushitem{\textit{#1} \qtybf{#2} = \qtybf{#3}} \nopagebreak\\[\parskip]} |
684 |
\def\opargbool#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]} |
|
43014 | 685 |
\def\opargboolorsmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{smart\_bool}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]} |
36926 | 686 |
|
687 |
Sledgehammer's options are categorized as follows:\ mode of operation |
|
38984 | 688 |
(\S\ref{mode-of-operation}), problem encoding (\S\ref{problem-encoding}), |
689 |
relevance filter (\S\ref{relevance-filter}), output format |
|
43038 | 690 |
(\S\ref{output-format}), authentication (\S\ref{authentication}), and timeouts |
691 |
(\S\ref{timeouts}). |
|
36926 | 692 |
|
693 |
The descriptions below refer to the following syntactic quantities: |
|
694 |
||
695 |
\begin{enum} |
|
696 |
\item[$\bullet$] \qtybf{string}: A string. |
|
697 |
\item[$\bullet$] \qtybf{bool\/}: \textit{true} or \textit{false}. |
|
43014 | 698 |
\item[$\bullet$] \qtybf{smart\_bool\/}: \textit{true}, \textit{false}, or |
40203 | 699 |
\textit{smart}. |
36926 | 700 |
\item[$\bullet$] \qtybf{int\/}: An integer. |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42722
diff
changeset
|
701 |
%\item[$\bullet$] \qtybf{float\/}: A floating-point number (e.g., 2.5). |
40343
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
702 |
\item[$\bullet$] \qtybf{float\_pair\/}: A pair of floating-point numbers |
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
703 |
(e.g., 0.6 0.95). |
43014 | 704 |
\item[$\bullet$] \qtybf{smart\_int\/}: An integer or \textit{smart}. |
43036 | 705 |
\item[$\bullet$] \qtybf{float\_or\_none\/}: A floating-point number (e.g., 60 or |
706 |
0.5) expressing a number of seconds, or the keyword \textit{none} ($\infty$ |
|
707 |
seconds). |
|
36926 | 708 |
\end{enum} |
709 |
||
43217 | 710 |
Default values are indicated in curly brackets (\textrm{\{\}}). Boolean options |
711 |
have a negated counterpart (e.g., \textit{blocking} vs.\ |
|
712 |
\textit{non\_blocking}). When setting them, ``= \textit{true}'' may be omitted. |
|
36926 | 713 |
|
714 |
\subsection{Mode of Operation} |
|
715 |
\label{mode-of-operation} |
|
716 |
||
717 |
\begin{enum} |
|
43014 | 718 |
\opnodefaultbrk{provers}{string} |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
719 |
Specifies the automatic provers to use as a space-separated list (e.g., |
43014 | 720 |
``\textit{e}~\textit{spass}~\textit{remote\_vampire}''). The following local |
721 |
provers are supported: |
|
36926 | 722 |
|
723 |
\begin{enum} |
|
42945 | 724 |
\item[$\bullet$] \textbf{\textit{cvc3}:} CVC3 is an SMT solver developed by |
725 |
Clark Barrett, Cesare Tinelli, and their colleagues \cite{cvc3}. To use CVC3, |
|
726 |
set the environment variable \texttt{CVC3\_SOLVER} to the complete path of the |
|
727 |
executable, including the file name. Sledgehammer has been tested with version |
|
728 |
2.2. |
|
729 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
730 |
\item[$\bullet$] \textbf{\textit{e}:} E is a first-order resolution prover |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
731 |
developed by Stephan Schulz \cite{schulz-2002}. To use E, set the environment |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
732 |
variable \texttt{E\_HOME} to the directory that contains the \texttt{eproof} |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
733 |
executable, or install the prebuilt E package from Isabelle's download page. See |
36926 | 734 |
\S\ref{installation} for details. |
735 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
736 |
\item[$\bullet$] \textbf{\textit{spass}:} SPASS is a first-order resolution |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
737 |
prover developed by Christoph Weidenbach et al.\ \cite{weidenbach-et-al-2009}. |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
738 |
To use SPASS, set the environment variable \texttt{SPASS\_HOME} to the directory |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
739 |
that contains the \texttt{SPASS} executable, or install the prebuilt SPASS |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
740 |
package from Isabelle's download page. Sledgehammer requires version 3.5 or |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
741 |
above. See \S\ref{installation} for details. |
36926 | 742 |
|
42945 | 743 |
\item[$\bullet$] \textbf{\textit{yices}:} Yices is an SMT solver developed at |
744 |
SRI \cite{yices}. To use Yices, set the environment variable |
|
745 |
\texttt{YICES\_SOLVER} to the complete path of the executable, including the |
|
746 |
file name. Sledgehammer has been tested with version 1.0. |
|
747 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
748 |
\item[$\bullet$] \textbf{\textit{vampire}:} Vampire is a first-order resolution |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
749 |
prover developed by Andrei Voronkov and his colleagues |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
750 |
\cite{riazanov-voronkov-2002}. To use Vampire, set the environment variable |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
751 |
\texttt{VAMPIRE\_HOME} to the directory that contains the \texttt{vampire} |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
752 |
executable. Sledgehammer has been tested with versions 11, 0.6, and 1.0. |
40942 | 753 |
|
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
754 |
\item[$\bullet$] \textbf{\textit{z3}:} Z3 is an SMT solver developed at |
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
755 |
Microsoft Research \cite{z3}. To use Z3, set the environment variable |
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
756 |
\texttt{Z3\_SOLVER} to the complete path of the executable, including the file |
42945 | 757 |
name, and set \texttt{Z3\_NON\_COMMERCIAL=yes} to confirm that you are a |
758 |
noncommercial user. Sledgehammer has been tested with versions 2.7 to 2.18. |
|
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
759 |
|
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
760 |
\item[$\bullet$] \textbf{\textit{z3\_atp}:} This version of Z3 pretends to be an |
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
761 |
ATP, exploiting Z3's undocumented support for the TPTP format. It is included |
42442 | 762 |
for experimental purposes. It requires version 2.18 or above. |
43053 | 763 |
|
764 |
\item[$\bullet$] \textbf{\textit{metis}:} Although it is much less powerful than |
|
765 |
the external provers, Metis itself can be used for proof search. |
|
766 |
||
767 |
\item[$\bullet$] \textbf{\textit{metisFT}:} Fully typed version of Metis. |
|
42945 | 768 |
\end{enum} |
769 |
||
770 |
In addition, the following remote provers are supported: |
|
771 |
||
772 |
\begin{enum} |
|
773 |
\item[$\bullet$] \textbf{\textit{remote\_cvc3}:} The remote version of CVC3 runs |
|
774 |
on servers at the TU M\"unchen (or wherever \texttt{REMOTE\_SMT\_URL} is set to |
|
775 |
point). |
|
40073 | 776 |
|
38601 | 777 |
\item[$\bullet$] \textbf{\textit{remote\_e}:} The remote version of E runs |
36926 | 778 |
on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}. |
779 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
780 |
\item[$\bullet$] \textbf{\textit{remote\_leo2}:} LEO-II is an automatic |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
781 |
higher-order prover developed by Christoph Benzm\"uller et al. \cite{leo2}. The |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
782 |
remote version of LEO-II runs on Geoff Sutcliffe's Miami servers. In the current |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
783 |
setup, the problems given to LEO-II are only mildly higher-order. |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
784 |
|
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
785 |
\item[$\bullet$] \textbf{\textit{remote\_satallax}:} Satallax is an automatic |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
786 |
higher-order prover developed by Chad Brown et al. \cite{satallax}. The remote |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
787 |
version of Satallax runs on Geoff Sutcliffe's Miami servers. In the current |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
788 |
setup, the problems given to Satallax are only mildly higher-order. |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
789 |
|
38601 | 790 |
\item[$\bullet$] \textbf{\textit{remote\_sine\_e}:} SInE-E is a metaprover |
791 |
developed by Kry\v stof Hoder \cite{sine} based on E. The remote version of |
|
792 |
SInE runs on Geoff Sutcliffe's Miami servers. |
|
793 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
794 |
\item[$\bullet$] \textbf{\textit{remote\_snark}:} SNARK is a first-order |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
795 |
resolution prover developed by Stickel et al.\ \cite{snark}. The remote version |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
796 |
of SNARK runs on Geoff Sutcliffe's Miami servers. |
40073 | 797 |
|
42945 | 798 |
\item[$\bullet$] \textbf{\textit{remote\_tofof\_e}:} ToFoF-E is a metaprover |
799 |
developed by Geoff Sutcliffe \cite{tofof} based on E running on his Miami |
|
800 |
servers. This ATP supports a fragment of the TPTP many-typed first-order format |
|
801 |
(TFF). It is supported primarily for experimenting with the |
|
802 |
\textit{type\_sys} $=$ \textit{simple} option (\S\ref{problem-encoding}). |
|
42940 | 803 |
|
42945 | 804 |
\item[$\bullet$] \textbf{\textit{remote\_vampire}:} The remote version of |
805 |
Vampire runs on Geoff Sutcliffe's Miami servers. Version 9 is used. |
|
806 |
||
807 |
\item[$\bullet$] \textbf{\textit{remote\_waldmeister}:} Waldmeister is a unit |
|
808 |
equality prover developed by Hillenbrand et al.\ \cite{waldmeister}. It can be |
|
809 |
used to prove universally quantified equations using unconditional equations. |
|
810 |
The remote version of Waldmeister runs on Geoff Sutcliffe's Miami servers. |
|
41738
eb98c60a6cf0
added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
blanchet
parents:
41727
diff
changeset
|
811 |
|
40942 | 812 |
\item[$\bullet$] \textbf{\textit{remote\_z3}:} The remote version of Z3 runs on |
813 |
servers at the TU M\"unchen (or wherever \texttt{REMOTE\_SMT\_URL} is set to |
|
814 |
point). |
|
40073 | 815 |
|
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
816 |
\item[$\bullet$] \textbf{\textit{remote\_z3\_atp}:} The remote version of ``Z3 |
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
817 |
as an ATP'' runs on Geoff Sutcliffe's Miami servers. |
36926 | 818 |
\end{enum} |
819 |
||
40942 | 820 |
By default, Sledgehammer will run E, SPASS, Vampire, SInE-E, and Z3 (or whatever |
42228 | 821 |
the SMT module's \textit{smt\_solver} configuration option is set to) in |
40073 | 822 |
parallel---either locally or remotely, depending on the number of processor |
823 |
cores available. For historical reasons, the default value of this option can be |
|
824 |
overridden using the option ``Sledgehammer: Provers'' from the ``Isabelle'' menu |
|
825 |
in Proof General. |
|
36926 | 826 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
827 |
It is a good idea to run several provers in parallel, although it could slow |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
828 |
down your machine. Running E, SPASS, and Vampire for 5~seconds yields a similar |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
829 |
success rate to running the most effective of these for 120~seconds |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
830 |
\cite{boehme-nipkow-2010}. |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
831 |
|
43053 | 832 |
For the \textit{min} subcommand, the default prover is \textit{metis}. If |
833 |
several provers are set, the first one is used. |
|
834 |
||
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
835 |
\opnodefault{prover}{string} |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
836 |
Alias for \textit{provers}. |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
837 |
|
42884 | 838 |
%\opnodefault{atps}{string} |
839 |
%Legacy alias for \textit{provers}. |
|
36926 | 840 |
|
42884 | 841 |
%\opnodefault{atp}{string} |
842 |
%Legacy alias for \textit{provers}. |
|
36926 | 843 |
|
38983 | 844 |
\opfalse{blocking}{non\_blocking} |
845 |
Specifies whether the \textbf{sledgehammer} command should operate |
|
846 |
synchronously. The asynchronous (non-blocking) mode lets the user start proving |
|
847 |
the putative theorem manually while Sledgehammer looks for a proof, but it can |
|
42995
e23f61546cf0
always run Sledgehammer synchronously in the jEdit interface (until the multithreading support for Proof General is ported)
blanchet
parents:
42964
diff
changeset
|
848 |
also be more confusing. Irrespective of the value of this option, Sledgehammer |
e23f61546cf0
always run Sledgehammer synchronously in the jEdit interface (until the multithreading support for Proof General is ported)
blanchet
parents:
42964
diff
changeset
|
849 |
is always run synchronously for the new jEdit-based user interface or if |
e23f61546cf0
always run Sledgehammer synchronously in the jEdit interface (until the multithreading support for Proof General is ported)
blanchet
parents:
42964
diff
changeset
|
850 |
\textit{debug} (\S\ref{output-format}) is enabled. |
38983 | 851 |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
852 |
\optrue{slicing}{no\_slicing} |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
853 |
Specifies whether the time allocated to a prover should be sliced into several |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
854 |
segments, each of which has its own set of possibly prover-dependent options. |
42446 | 855 |
For SPASS and Vampire, the first slice tries the fast but incomplete |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
856 |
set-of-support (SOS) strategy, whereas the second slice runs without it. For E, |
42446 | 857 |
up to three slices are tried, with different weighted search strategies and |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
858 |
number of facts. For SMT solvers, several slices are tried with the same options |
42446 | 859 |
each time but fewer and fewer facts. According to benchmarks with a timeout of |
860 |
30 seconds, slicing is a valuable optimization, and you should probably leave it |
|
861 |
enabled unless you are conducting experiments. This option is implicitly |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
862 |
disabled for (short) automatic runs. |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
863 |
|
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
864 |
\nopagebreak |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
865 |
{\small See also \textit{verbose} (\S\ref{output-format}).} |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42442
diff
changeset
|
866 |
|
36926 | 867 |
\opfalse{overlord}{no\_overlord} |
868 |
Specifies whether Sledgehammer should put its temporary files in |
|
869 |
\texttt{\$ISA\-BELLE\_\allowbreak HOME\_\allowbreak USER}, which is useful for |
|
870 |
debugging Sledgehammer but also unsafe if several instances of the tool are run |
|
871 |
simultaneously. The files are identified by the prefix \texttt{prob\_}; you may |
|
872 |
safely remove them after Sledgehammer has run. |
|
873 |
||
874 |
\nopagebreak |
|
875 |
{\small See also \textit{debug} (\S\ref{output-format}).} |
|
876 |
\end{enum} |
|
877 |
||
878 |
\subsection{Problem Encoding} |
|
879 |
\label{problem-encoding} |
|
880 |
||
881 |
\begin{enum} |
|
882 |
\opfalse{full\_types}{partial\_types} |
|
43216 | 883 |
Specifies whether a type-sound encoding should be used when translating |
884 |
higher-order types to untyped first-order logic. Enabling this option virtually |
|
885 |
prevents the discovery of type-incorrect proofs, but it can slow down the ATP |
|
886 |
slightly. This option is implicitly enabled for automatic runs. For historical |
|
887 |
reasons, the default value of this option can be overridden using the option |
|
888 |
``Sledgehammer: Full Types'' from the ``Isabelle'' menu in Proof General. |
|
42228 | 889 |
|
890 |
\opdefault{type\_sys}{string}{smart} |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
891 |
Specifies the type system to use in ATP problems. Some of the type systems are |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
892 |
unsound, meaning that they can give rise to spurious proofs (unreconstructible |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
893 |
using Metis). The supported type systems are listed below, with an indication of |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
894 |
their soundness in parentheses: |
42228 | 895 |
|
896 |
\begin{enum} |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
897 |
\item[$\bullet$] \textbf{\textit{erased} (very unsound):} No type information is |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
898 |
supplied to the ATP. Types are simply erased. |
42582 | 899 |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
900 |
\item[$\bullet$] \textbf{\textit{poly\_preds} (sound):} Types are encoded using |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
901 |
a predicate \textit{has\_\allowbreak type\/}$(\tau, t)$ that restricts the range |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
902 |
of bound variables. Constants are annotated with their types, supplied as extra |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
903 |
arguments, to resolve overloading. |
42685 | 904 |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
905 |
\item[$\bullet$] \textbf{\textit{poly\_tags} (sound):} Each term and subterm is |
43002
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
906 |
tagged with its type using a function $\mathit{type\_info\/}(\tau, t)$. This |
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
907 |
coincides with the encoding used by the \textit{metisFT} command. |
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
908 |
|
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
909 |
\item[$\bullet$] \textbf{\textit{poly\_args} (unsound):} |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
910 |
Like for \textit{poly\_preds} constants are annotated with their types to |
43002
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
911 |
resolve overloading, but otherwise no type information is encoded. This |
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
912 |
coincides with the encoding used by the \textit{metis} command (before it falls |
43203
6c8170f8849e
don't mention "metisX" so much in the docs -- it will go away soon
blanchet
parents:
43170
diff
changeset
|
913 |
back on \textit{metisFT}). |
42685 | 914 |
|
42722 | 915 |
\item[$\bullet$] |
916 |
\textbf{% |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
917 |
\textit{mono\_preds}, \textit{mono\_tags} (sound); |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
918 |
\textit{mono\_args} (unsound):} \\ |
42722 | 919 |
Similar to \textit{poly\_preds}, \textit{poly\_tags}, and \textit{poly\_args}, |
920 |
respectively, but the problem is additionally monomorphized, meaning that type |
|
921 |
variables are instantiated with heuristically chosen ground types. |
|
922 |
Monomorphization can simplify reasoning but also leads to larger fact bases, |
|
923 |
which can slow down the ATPs. |
|
42582 | 924 |
|
42722 | 925 |
\item[$\bullet$] |
926 |
\textbf{% |
|
927 |
\textit{mangled\_preds}, |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
928 |
\textit{mangled\_tags} (sound); \\ |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
929 |
\textit{mangled\_args} (unsound):} \\ |
42722 | 930 |
Similar to |
931 |
\textit{mono\_preds}, \textit{mono\_tags}, and \textit{mono\_args}, |
|
932 |
respectively but types are mangled in constant names instead of being supplied |
|
933 |
as ground term arguments. The binary predicate $\mathit{has\_type\/}(\tau, t)$ |
|
934 |
becomes a unary predicate $\mathit{has\_type\_}\tau(t)$, and the binary function |
|
42589
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42582
diff
changeset
|
935 |
$\mathit{type\_info\/}(\tau, t)$ becomes a unary function |
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42582
diff
changeset
|
936 |
$\mathit{type\_info\_}\tau(t)$. |
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42582
diff
changeset
|
937 |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
938 |
\item[$\bullet$] \textbf{\textit{simple} (sound):} Use the prover's support for |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
939 |
simple types if available; otherwise, fall back on \textit{mangled\_preds}. The |
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
940 |
problem is monomorphized. |
42681 | 941 |
|
942 |
\item[$\bullet$] |
|
943 |
\textbf{% |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
944 |
\textit{poly\_preds}?, \textit{poly\_tags}?, \textit{mono\_preds}?, \textit{mono\_tags}?, \\ |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
945 |
\textit{mangled\_preds}?, \textit{mangled\_tags}?, \textit{simple}? (quasi-sound):} \\ |
42743
b81127eb79f3
reflect option renaming in doc + do not document the type systems poly_preds? and poly_tags?, since they are virtually identical to the non-? versions
blanchet
parents:
42736
diff
changeset
|
946 |
The type systems \textit{poly\_preds}, \textit{poly\_tags}, |
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
947 |
\textit{mono\_preds}, \textit{mono\_tags}, \textit{mangled\_preds}, |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
948 |
\textit{mangled\_tags}, and \textit{simple} are fully typed and sound. For each |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
949 |
of these, Sledgehammer also provides a lighter, virtually sound variant |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
950 |
identified by a question mark (`{?}')\ that detects and erases monotonic types, |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
951 |
notably infinite types. (For \textit{simple}, the types are not actually erased |
43216 | 952 |
but rather replaced by a shared uniform type of individuals.) As argument to the |
953 |
\textit{metis} proof method, the question mark is replaced by a |
|
954 |
``\textit{\_query}'' suffix. |
|
42582 | 955 |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
956 |
\item[$\bullet$] |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
957 |
\textbf{% |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
958 |
\textit{poly\_preds}!, \textit{poly\_tags}!, \textit{mono\_preds}!, \textit{mono\_tags}!, \\ |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
959 |
\textit{mangled\_preds}!, \textit{mangled\_tags}!, \textit{simple}! \\ |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
960 |
(mildly unsound):} \\ |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
961 |
The type systems \textit{poly\_preds}, \textit{poly\_tags}, |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
962 |
\textit{mono\_preds}, \textit{mono\_tags}, \textit{mangled\_preds}, |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
963 |
\textit{mangled\_tags}, and \textit{simple} also admit a mildly unsound (but |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
964 |
very efficient) variant identified by an exclamation mark (`{!}') that detects |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
965 |
and erases erases all types except those that are clearly finite (e.g., |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
966 |
\textit{bool}). (For \textit{simple}, the types are not actually erased but |
43216 | 967 |
rather replaced by a shared uniform type of individuals.) As argument to the |
968 |
\textit{metis} proof method, the exclamation mark is replaced by a |
|
969 |
``\textit{\_bang}'' suffix. |
|
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
970 |
|
42237 | 971 |
\item[$\bullet$] \textbf{\textit{smart}:} If \textit{full\_types} is enabled, |
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
972 |
uses a sound or virtually sound encoding; otherwise, uses any encoding. The actual |
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
973 |
encoding used depends on the ATP and should be the most efficient for that ATP. |
42228 | 974 |
\end{enum} |
975 |
||
42856 | 976 |
In addition, all the \textit{preds} and \textit{tags} type systems are available |
977 |
in two variants, a lightweight and a heavyweight variant. The lightweight |
|
978 |
variants are generally more efficient and are the default; the heavyweight |
|
979 |
variants are identified by a \textit{\_heavy} suffix (e.g., |
|
980 |
\textit{mangled\_preds\_heavy}{?}). |
|
42523
08346ea46a59
added (without implementation yet) new type encodings for Sledgehammer/ATP
blanchet
parents:
42511
diff
changeset
|
981 |
|
42856 | 982 |
For SMT solvers and ToFoF-E, the type system is always \textit{simple}, |
983 |
irrespective of the value of this option. |
|
42888 | 984 |
|
985 |
\nopagebreak |
|
986 |
{\small See also \textit{max\_new\_mono\_instances} (\S\ref{relevance-filter}) |
|
987 |
and \textit{max\_mono\_iters} (\S\ref{relevance-filter}).} |
|
43065 | 988 |
|
989 |
\opsmart{explicit\_apply}{implicit\_apply} |
|
990 |
Specifies whether function application should be encoded as an explicit |
|
991 |
``apply'' operator in ATP problems. If the option is set to \textit{false}, each |
|
992 |
function will be directly applied to as many arguments as possible. Disabling |
|
993 |
this option can sometimes prevent the discovery of higher-order proofs. |
|
38591 | 994 |
\end{enum} |
36926 | 995 |
|
38591 | 996 |
\subsection{Relevance Filter} |
997 |
\label{relevance-filter} |
|
998 |
||
999 |
\begin{enum} |
|
40343
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
1000 |
\opdefault{relevance\_thresholds}{float\_pair}{\upshape 0.45~0.85} |
38746 | 1001 |
Specifies the thresholds above which facts are considered relevant by the |
1002 |
relevance filter. The first threshold is used for the first iteration of the |
|
1003 |
relevance filter and the second threshold is used for the last iteration (if it |
|
1004 |
is reached). The effective threshold is quadratically interpolated for the other |
|
40343
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
1005 |
iterations. Each threshold ranges from 0 to 1, where 0 means that all theorems |
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
1006 |
are relevant and 1 only theorems that refer to previously seen constants. |
36926 | 1007 |
|
43065 | 1008 |
\opdefault{max\_relevant}{smart\_int}{smart} |
38746 | 1009 |
Specifies the maximum number of facts that may be returned by the relevance |
1010 |
filter. If the option is set to \textit{smart}, it is set to a value that was |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
1011 |
empirically found to be appropriate for the prover. A typical value would be |
43065 | 1012 |
250. |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41747
diff
changeset
|
1013 |
|
42884 | 1014 |
\opdefault{max\_new\_mono\_instances}{int}{\upshape 400} |
1015 |
Specifies the maximum number of monomorphic instances to generate beyond |
|
1016 |
\textit{max\_relevant}. The higher this limit is, the more monomorphic instances |
|
1017 |
are potentially generated. Whether monomorphization takes place depends on the |
|
1018 |
type system used. |
|
1019 |
||
1020 |
\nopagebreak |
|
1021 |
{\small See also \textit{type\_sys} (\S\ref{problem-encoding}).} |
|
1022 |
||
1023 |
\opdefault{max\_mono\_iters}{int}{\upshape 3} |
|
1024 |
Specifies the maximum number of iterations for the monomorphization fixpoint |
|
1025 |
construction. The higher this limit is, the more monomorphic instances are |
|
1026 |
potentially generated. Whether monomorphization takes place depends on the |
|
1027 |
type system used. |
|
1028 |
||
1029 |
\nopagebreak |
|
1030 |
{\small See also \textit{type\_sys} (\S\ref{problem-encoding}).} |
|
36926 | 1031 |
\end{enum} |
1032 |
||
1033 |
\subsection{Output Format} |
|
1034 |
\label{output-format} |
|
1035 |
||
1036 |
\begin{enum} |
|
1037 |
||
1038 |
\opfalse{verbose}{quiet} |
|
1039 |
Specifies whether the \textbf{sledgehammer} command should explain what it does. |
|
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
40942
diff
changeset
|
1040 |
This option is implicitly disabled for automatic runs. |
36926 | 1041 |
|
1042 |
\opfalse{debug}{no\_debug} |
|
40203 | 1043 |
Specifies whether Sledgehammer should display additional debugging information |
1044 |
beyond what \textit{verbose} already displays. Enabling \textit{debug} also |
|
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
40942
diff
changeset
|
1045 |
enables \textit{verbose} and \textit{blocking} (\S\ref{mode-of-operation}) |
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
40942
diff
changeset
|
1046 |
behind the scenes. The \textit{debug} option is implicitly disabled for |
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
40942
diff
changeset
|
1047 |
automatic runs. |
36926 | 1048 |
|
1049 |
\nopagebreak |
|
1050 |
{\small See also \textit{overlord} (\S\ref{mode-of-operation}).} |
|
1051 |
||
1052 |
\opfalse{isar\_proof}{no\_isar\_proof} |
|
1053 |
Specifies whether Isar proofs should be output in addition to one-liner |
|
1054 |
\textit{metis} proofs. Isar proof construction is still experimental and often |
|
1055 |
fails; however, they are usually faster and sometimes more robust than |
|
1056 |
\textit{metis} proofs. |
|
1057 |
||
40343
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
1058 |
\opdefault{isar\_shrink\_factor}{int}{\upshape 1} |
36926 | 1059 |
Specifies the granularity of the Isar proof. A value of $n$ indicates that each |
1060 |
Isar proof step should correspond to a group of up to $n$ consecutive proof |
|
1061 |
steps in the ATP proof. |
|
1062 |
\end{enum} |
|
1063 |
||
38984 | 1064 |
\subsection{Authentication} |
1065 |
\label{authentication} |
|
1066 |
||
1067 |
\begin{enum} |
|
1068 |
\opnodefault{expect}{string} |
|
1069 |
Specifies the expected outcome, which must be one of the following: |
|
36926 | 1070 |
|
1071 |
\begin{enum} |
|
40203 | 1072 |
\item[$\bullet$] \textbf{\textit{some}:} Sledgehammer found a (potentially |
1073 |
unsound) proof. |
|
38984 | 1074 |
\item[$\bullet$] \textbf{\textit{none}:} Sledgehammer found no proof. |
43014 | 1075 |
\item[$\bullet$] \textbf{\textit{timeout}:} Sledgehammer timed out. |
40203 | 1076 |
\item[$\bullet$] \textbf{\textit{unknown}:} Sledgehammer encountered some |
1077 |
problem. |
|
38984 | 1078 |
\end{enum} |
1079 |
||
1080 |
Sledgehammer emits an error (if \textit{blocking} is enabled) or a warning |
|
1081 |
(otherwise) if the actual outcome differs from the expected outcome. This option |
|
1082 |
is useful for regression testing. |
|
1083 |
||
1084 |
\nopagebreak |
|
43038 | 1085 |
{\small See also \textit{blocking} (\S\ref{mode-of-operation}) and |
1086 |
\textit{timeout} (\S\ref{timeouts}).} |
|
1087 |
\end{enum} |
|
1088 |
||
1089 |
\subsection{Timeouts} |
|
1090 |
\label{timeouts} |
|
1091 |
||
1092 |
\begin{enum} |
|
1093 |
\opdefault{timeout}{float\_or\_none}{\upshape 30} |
|
1094 |
Specifies the maximum number of seconds that the automatic provers should spend |
|
1095 |
searching for a proof. This excludes problem preparation and is a soft limit. |
|
1096 |
For historical reasons, the default value of this option can be overridden using |
|
1097 |
the option ``Sledgehammer: Time Limit'' from the ``Isabelle'' menu in Proof |
|
1098 |
General. |
|
1099 |
||
1100 |
\opdefault{preplay\_timeout}{float\_or\_none}{\upshape 4} |
|
1101 |
Specifies the maximum number of seconds that Metis should be spent trying to |
|
1102 |
``preplay'' the found proof. If this option is set to 0, no preplaying takes |
|
1103 |
place, and no timing information is displayed next to the suggested Metis calls. |
|
36926 | 1104 |
\end{enum} |
1105 |
||
1106 |
\let\em=\sl |
|
1107 |
\bibliography{../manual}{} |
|
1108 |
\bibliographystyle{abbrv} |
|
1109 |
||
1110 |
\end{document} |