author | Lukas Bartl |
Thu, 09 Jan 2025 13:18:37 +0100 | |
changeset 81757 | 4d15005da582 |
parent 81613 | 44afa6f1baad |
child 82208 | bab8158a02f0 |
permissions | -rw-r--r-- |
36926 | 1 |
\documentclass[a4paper,12pt]{article} |
2 |
\usepackage[T1]{fontenc} |
|
3 |
\usepackage{amsmath} |
|
4 |
\usepackage{amssymb} |
|
5 |
\usepackage{color} |
|
6 |
\usepackage{footmisc} |
|
7 |
\usepackage{graphicx} |
|
8 |
%\usepackage{mathpazo} |
|
9 |
\usepackage{multicol} |
|
10 |
\usepackage{stmaryrd} |
|
11 |
%\usepackage[scaled=.85]{beramono} |
|
48962
a1acc1cb0271
more standard document preparation within session context;
wenzelm
parents:
48803
diff
changeset
|
12 |
\usepackage{isabelle,iman,pdfsetup} |
36926 | 13 |
|
68649 | 14 |
\newcommand\download{\url{https://isabelle.in.tum.de/components/}} |
46242 | 15 |
|
57040 | 16 |
\let\oldS=\S |
17 |
\def\S{\oldS\,} |
|
18 |
||
43216 | 19 |
\def\qty#1{\ensuremath{\left<\mathit{#1\/}\right>}} |
20 |
\def\qtybf#1{$\mathbf{\left<\textbf{\textit{#1\/}}\right>}$} |
|
21 |
||
45516 | 22 |
\newcommand\const[1]{\textsf{#1}} |
23 |
||
36926 | 24 |
%\oddsidemargin=4.6mm |
25 |
%\evensidemargin=4.6mm |
|
26 |
%\textwidth=150mm |
|
27 |
%\topmargin=4.6mm |
|
28 |
%\headheight=0mm |
|
29 |
%\headsep=0mm |
|
30 |
%\textheight=234mm |
|
31 |
||
32 |
\def\Colon{\mathord{:\mkern-1.5mu:}} |
|
33 |
%\def\lbrakk{\mathopen{\lbrack\mkern-3.25mu\lbrack}} |
|
34 |
%\def\rbrakk{\mathclose{\rbrack\mkern-3.255mu\rbrack}} |
|
35 |
\def\lparr{\mathopen{(\mkern-4mu\mid}} |
|
36 |
\def\rparr{\mathclose{\mid\mkern-4mu)}} |
|
37 |
||
38 |
\def\unk{{?}} |
|
80091
36389d25d33e
rename \undef to \undefined to avoid problems with MacTeX 2014 or Ubuntu 24.04 beta;
wenzelm
parents:
78149
diff
changeset
|
39 |
\def\undefined{(\lambda x.\; \unk)} |
36926 | 40 |
%\def\unr{\textit{others}} |
41 |
\def\unr{\ldots} |
|
73595
aece5cc9efb7
simplified typesetting of \<guillemotleft>...\<guillemotright>;
wenzelm
parents:
73436
diff
changeset
|
42 |
\def\Abs#1{\hbox{\rm{\guillemetleft}}{\,#1\,}\hbox{\rm{\guillemetright}}} |
36926 | 43 |
\def\Q{{\smash{\lower.2ex\hbox{$\scriptstyle?$}}}} |
44 |
||
45 |
\urlstyle{tt} |
|
46 |
||
55290 | 47 |
\renewcommand\_{\hbox{\textunderscore\kern-.05ex}} |
48 |
||
72589 | 49 |
\hyphenation{Isa-belle super-posi-tion zipper-posi-tion} |
50 |
||
36926 | 51 |
\begin{document} |
52 |
||
45516 | 53 |
%%% TYPESETTING |
54 |
%\renewcommand\labelitemi{$\bullet$} |
|
55 |
\renewcommand\labelitemi{\raise.065ex\hbox{\small\textbullet}} |
|
56 |
||
73723 | 57 |
\title{\includegraphics[scale=0.5]{isabelle_logo} \\[4ex] |
36926 | 58 |
Hammering Away \\[\smallskipamount] |
59 |
\Large A User's Guide to Sledgehammer for Isabelle/HOL} |
|
60 |
\author{\hbox{} \\ |
|
70818
13d6b561b0ea
added para constrasting 'primrec' and 'fun' -- and removed my middle name
blanchet
parents:
68649
diff
changeset
|
61 |
Jasmin Blanchette \\ |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
62 |
{\normalsize Institut f\"ur Informatik, Ludwig-Maximilians-Universit\"at M\"unchen} \\[4\smallskipamount] |
43002
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
63 |
{\normalsize with contributions from} \\[4\smallskipamount] |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
64 |
Martin Desharnais \\ |
81612
6528e378be87
updated affiliation in Sledgehammer documentation
desharna
parents:
81286
diff
changeset
|
65 |
{\normalsize Max-Planck-Institut f\"ur Informatik} \\[4\smallskipamount] |
43002
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
66 |
Lawrence C. Paulson \\ |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
67 |
{\normalsize Computer Laboratory, University of Cambridge} \\[4\smallskipamount] |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
68 |
Lukas Bartl \\ |
81757
4d15005da582
tuned documentation and order of instantiated facts
Lukas Bartl
parents:
81613
diff
changeset
|
69 |
{\normalsize Institut f\"ur Informatik, Universit\"at Augsburg} \\ |
36926 | 70 |
\hbox{}} |
71 |
||
72 |
\maketitle |
|
73 |
||
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
74 |
\newpage |
36926 | 75 |
\tableofcontents |
76 |
||
77 |
\setlength{\parskip}{.7em plus .2em minus .1em} |
|
78 |
\setlength{\parindent}{0pt} |
|
79 |
\setlength{\abovedisplayskip}{\parskip} |
|
80 |
\setlength{\abovedisplayshortskip}{.9\parskip} |
|
81 |
\setlength{\belowdisplayskip}{\parskip} |
|
82 |
\setlength{\belowdisplayshortskip}{.9\parskip} |
|
83 |
||
52078 | 84 |
% general-purpose enum environment with correct spacing |
36926 | 85 |
\newenvironment{enum}% |
86 |
{\begin{list}{}{% |
|
87 |
\setlength{\topsep}{.1\parskip}% |
|
88 |
\setlength{\partopsep}{.1\parskip}% |
|
89 |
\setlength{\itemsep}{\parskip}% |
|
90 |
\advance\itemsep by-\parsep}} |
|
91 |
{\end{list}} |
|
92 |
||
93 |
\def\pre{\begingroup\vskip0pt plus1ex\advance\leftskip by\leftmargin |
|
94 |
\advance\rightskip by\leftmargin} |
|
95 |
\def\post{\vskip0pt plus1ex\endgroup} |
|
96 |
||
97 |
\def\prew{\pre\advance\rightskip by-\leftmargin} |
|
98 |
\def\postw{\post} |
|
99 |
||
68565 | 100 |
|
36926 | 101 |
\section{Introduction} |
102 |
\label{introduction} |
|
103 |
||
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
104 |
Sledgehammer is a tool that applies automatic theorem provers (ATPs) |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
105 |
and satisfiability-modulo-theories (SMT) solvers on the current goal, mostly |
77428 | 106 |
to find proofs but also to refute the goal.% |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
107 |
\footnote{The distinction between ATPs and SMT solvers is mostly |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
108 |
historical but convenient.} |
47561
92d88c89efff
update documentation (mostly based on feedback by Makarius)
blanchet
parents:
47530
diff
changeset
|
109 |
% |
72403
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
110 |
The supported ATPs include agsyHOL \cite{agsyHOL}, Alt-Ergo \cite{alt-ergo}, E |
72592 | 111 |
\cite{schulz-2019}, iProver \cite{korovin-2009}, LEO-II \cite{leo2}, Leo-III |
72403
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
112 |
\cite{leo3}, Satallax \cite{satallax}, SPASS \cite{weidenbach-et-al-2009}, |
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
113 |
Vampire \cite{riazanov-voronkov-2002}, Waldmeister \cite{waldmeister}, and |
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
114 |
Zipperposition \cite{cruanes-2014}. The ATPs are run either locally or remotely |
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
115 |
via the System\-On\-TPTP web service \cite{sutcliffe-2000}. The supported SMT |
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
116 |
solvers are CVC4 \cite{cvc4}, cvc5 \cite{barbosa-et-al-cvc5}, veriT |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
117 |
\cite{bouton-et-al-2009}, and Z3 \cite{de-moura-2008}. These are always run |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
118 |
locally. |
36926 | 119 |
|
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
120 |
The problem passed to the automatic provers (or solvers) consists of your |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
121 |
current goal together with a heuristic selection of hundreds of facts (theorems) |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
122 |
from the current theory context, filtered by relevance. |
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
123 |
|
72403
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
124 |
The result of a successful proof search is some source text that typically |
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
125 |
reconstructs the proof within Isabelle. For ATPs, the reconstructed proof |
4a3169d8885c
removed support for obsolete prover SNARK and underperforming prover E-Par
blanchet
parents:
72402
diff
changeset
|
126 |
typically relies on the general-purpose \textit{metis} proof method, which |
45380 | 127 |
integrates the Metis ATP in Isabelle/HOL with explicit inferences going through |
128 |
the kernel. Thus its results are correct by construction. |
|
36926 | 129 |
|
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
130 |
Sometimes the automatic provers might detect that the goal is inconsistent with |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
131 |
the background facts---or even that the background facts are inconsistent |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
132 |
regardless of of the goal. |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
133 |
|
53760
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
134 |
For Isabelle/jEdit users, Sledgehammer provides an automatic mode that can be |
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
135 |
enabled via the ``Auto Sledgehammer'' option under ``Plugins > Plugin Options > |
54114 | 136 |
Isabelle > General.'' In this mode, a reduced version of Sledgehammer is run on |
137 |
every newly entered theorem for a few seconds. |
|
39320 | 138 |
|
36926 | 139 |
\newbox\boxA |
46298 | 140 |
\setbox\boxA=\hbox{\texttt{NOSPAM}} |
36926 | 141 |
|
72589 | 142 |
\newcommand\authoremail{\texttt{jasmin.blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak |
75021 | 143 |
gmail.\allowbreak com}} |
42763 | 144 |
|
40689 | 145 |
To run Sledgehammer, you must make sure that the theory \textit{Sledgehammer} is |
146 |
imported---this is rarely a problem in practice since it is part of |
|
72589 | 147 |
\textit{Main}. Examples of Sledgehammer use can be found in the |
148 |
\texttt{src/HOL/Metis\_Examples} directory. Comments and bug reports |
|
149 |
concerning Sledgehammer or this manual should be directed to the author at |
|
150 |
\authoremail. |
|
36926 | 151 |
|
152 |
||
153 |
\section{Installation} |
|
154 |
\label{installation} |
|
155 |
||
48387 | 156 |
Sledgehammer is part of Isabelle, so you do not need to install it. However, it |
46242 | 157 |
relies on third-party automatic provers (ATPs and SMT solvers). |
42763 | 158 |
|
74048 | 159 |
Among the ATPs, agsyHOL, Alt-Ergo, E, LEO-II, Leo-III, Satallax, SPASS, Vampire, |
160 |
and Zipperposition can be run locally; in addition, agsyHOL, Alt-Ergo, E, |
|
161 |
iProver, LEO-II, Leo-III, Satallax, Vampire, Waldmeister, and Zipperposition are |
|
162 |
available remotely via System\-On\-TPTP \cite{sutcliffe-2000}. The SMT solvers |
|
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
163 |
CVC4, cvc5, veriT, and Z3 can be run locally. |
36926 | 164 |
|
46242 | 165 |
There are three main ways to install automatic provers on your machine: |
36926 | 166 |
|
46242 | 167 |
\begin{sloppy} |
168 |
\begin{enum} |
|
169 |
\item[\labelitemi] If you installed an official Isabelle package, it should |
|
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
170 |
already include properly set up executables for CVC4, cvc5, E, SPASS, Vampire, |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
171 |
veriT, Z3, and Zipperposition ready to use. |
46242 | 172 |
|
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
173 |
\item[\labelitemi] Alternatively, you can download the Isabelle-aware CVC4, |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
174 |
cvc5, E, SPASS, Vampire, veriT, Z3, and Zipperposition binary packages from |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
175 |
\download. Extract the archives, then add a line to your |
75021 | 176 |
\texttt{\$ISABELLE\_HOME\_USER\slash etc\slash components}% |
41747
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
177 |
\footnote{The variable \texttt{\$ISABELLE\_HOME\_USER} is set by Isabelle at |
46242 | 178 |
startup. Its value can be retrieved by executing \texttt{isabelle} |
41747
f58d4d202924
fix path to etc/settings and etc/components in doc
blanchet
parents:
41740
diff
changeset
|
179 |
\texttt{getenv} \texttt{ISABELLE\_HOME\_USER} on the command line.} |
72592 | 180 |
file with the absolute path to the prover. For example, if the |
181 |
\texttt{components} file does not exist yet and you extracted SPASS to |
|
182 |
\texttt{/usr/local/spass-3.8ds}, create it with the single line |
|
36926 | 183 |
|
184 |
\prew |
|
47577 | 185 |
\texttt{/usr/local/spass-3.8ds} |
36926 | 186 |
\postw |
187 |
||
47561
92d88c89efff
update documentation (mostly based on feedback by Makarius)
blanchet
parents:
47530
diff
changeset
|
188 |
in it. |
38043 | 189 |
|
70937 | 190 |
\item[\labelitemi] If you prefer to build agsyHOL, Alt-Ergo, E, LEO-II, |
75021 | 191 |
Leo-III, Satallax, or Zipperposition manually, set the environment variable |
52078 | 192 |
\texttt{AGSYHOL\_HOME}, \texttt{E\_HOME}, \texttt{LEO2\_HOME}, |
75021 | 193 |
\texttt{LEO3\_HOME}, \texttt{SATALLAX\_HOME}, or \texttt{ZIPPERPOSITION\_HOME} |
68565 | 194 |
to the directory that contains the \texttt{agsyHOL}, |
75021 | 195 |
\texttt{eprover} (or \texttt{eprover-ho}), |
196 |
\texttt{leo}, \texttt{leo3}, \texttt{satallax}, or \texttt{zipperposition} |
|
197 |
executable; for Alt-Ergo, set the environment variable \texttt{WHY3\_HOME} to the |
|
73970
34c8cf767fa3
adjusted E setup to avoid generating FOOL with 2.5 (where 'ite' is missing)
blanchet
parents:
73941
diff
changeset
|
198 |
directory that contains the \texttt{why3} executable. Ideally, you |
70929 | 199 |
should also set \texttt{E\_VERSION}, \texttt{LEO2\_VERSION}, |
75021 | 200 |
\texttt{LEO3\_VERSION}, \texttt{SATALLAX\_VERSION}, or |
201 |
\texttt{ZIPPERPOSITION\_VERSION} to the prover's version number (e.g., ``3.6''). |
|
36926 | 202 |
|
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
203 |
Similarly, if you want to install CVC4, cvc5, veriT, or Z3, set the environment |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
204 |
variable \texttt{CVC4\_\allowbreak SOLVER}, \texttt{CVC5\_\allowbreak SOLVER}, |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
205 |
\texttt{ISABELLE\_\allowbreak VERIT}, |
74048 | 206 |
or \texttt{Z3\_SOLVER} to the complete path of the executable, \emph{including |
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
207 |
the file name}. Ideally, also set \texttt{CVC4\_VERSION}, \texttt{CVC5\_VERSION}, |
74048 | 208 |
\texttt{VERIT\_VERSION}, or \texttt{Z3\_VERSION} to the solver's version number |
209 |
(e.g., ``4.4.0''). |
|
46242 | 210 |
\end{enum} |
211 |
\end{sloppy} |
|
36926 | 212 |
|
66735 | 213 |
To check whether the provers are successfully installed, try out the example |
214 |
in \S\ref{first-steps}. If the remote versions of any of these provers is used |
|
215 |
(identified by the prefix ``\textit{remote\_\/}''), or if the local versions |
|
216 |
fail to solve the easy goal presented there, something must be wrong with the |
|
217 |
installation. |
|
46242 | 218 |
|
68565 | 219 |
|
36926 | 220 |
\section{First Steps} |
221 |
\label{first-steps} |
|
222 |
||
223 |
To illustrate Sledgehammer in context, let us start a theory file and |
|
224 |
attempt to prove a simple lemma: |
|
225 |
||
226 |
\prew |
|
227 |
\textbf{theory}~\textit{Scratch} \\ |
|
75021 | 228 |
\noindent\hbox{}\quad \textbf{imports}~\textit{Main} \\ |
36926 | 229 |
\textbf{begin} \\[2\smallskipamount] |
230 |
% |
|
42945 | 231 |
\textbf{lemma} ``$[a] = [b] \,\Longrightarrow\, a = b$'' \\ |
36926 | 232 |
\textbf{sledgehammer} |
233 |
\postw |
|
234 |
||
53760
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
235 |
Instead of issuing the \textbf{sledgehammer} command, you can also use the |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
236 |
Sledgehammer panel in Isabelle/jEdit. Either way, Sledgehammer will produce |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
237 |
something like the following output after a few seconds: |
36926 | 238 |
|
239 |
\prew |
|
240 |
\slshape |
|
75021 | 241 |
e found a proof\ldots \\ |
242 |
cvc4 found a proof\ldots \\ |
|
243 |
z3 found a proof\ldots \\ |
|
244 |
vampire found a proof\ldots \\ |
|
245 |
e: Try this: \textbf{by} \textit{simp} (0.3 ms) \\ |
|
246 |
cvc4: Try this: \textbf{by} \textit{simp} (0.4 ms) \\ |
|
247 |
z3: Try this: \textbf{by} \textit{simp} (0.5 ms) \\ |
|
77422 | 248 |
vampire: Try this: \textbf{by} \textit{simp} (0.3 ms) |
36926 | 249 |
\postw |
250 |
||
75021 | 251 |
Sledgehammer ran CVC4, E, Vampire, Z3, and possibly other provers in parallel. |
252 |
The list may vary depending on which provers are installed and how many |
|
253 |
processor cores are available. |
|
36926 | 254 |
|
72592 | 255 |
For each successful prover, Sledgehammer gives a one-line Isabelle proof. Rough |
256 |
timings are shown in parentheses, indicating how fast the call is. You can |
|
257 |
click the proof to insert it into the theory text. |
|
36926 | 258 |
|
51190
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents:
51189
diff
changeset
|
259 |
In addition, you can ask Sledgehammer for an Isar text proof by enabling the |
49919 | 260 |
\textit{isar\_proofs} option (\S\ref{output-format}): |
36926 | 261 |
|
262 |
\prew |
|
49919 | 263 |
\textbf{sledgehammer} [\textit{isar\_proofs}] |
36926 | 264 |
\postw |
265 |
||
266 |
When Isar proof construction is successful, it can yield proofs that are more |
|
72592 | 267 |
readable and also faster than \textit{metis} or \textit{smt} one-line |
268 |
proofs. This feature is experimental. |
|
36926 | 269 |
|
78149
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
270 |
%For goals that are inconsistent with the background theory (and hence unprovable |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
271 |
%unless the theory is itself inconsistent), such as |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
272 |
% |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
273 |
%\prew |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
274 |
%\textbf{lemma} ``$\mathit{length}\; (\mathit{zip}\; \mathit{xs}\; \mathit{ys}) = \mathit{length}\; \mathit{xs}$'' \\ |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
275 |
%\textbf{sledgehammer} |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
276 |
%\postw |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
277 |
% |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
278 |
%Sledgehammer's output might look as follows: |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
279 |
% |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
280 |
%\prew |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
281 |
%\slshape |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
282 |
%vampire found a falsification\ldots \\ |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
283 |
%vampire: The goal is falsified by these facts: append\_Cons, nth\_append\_length, self\_append\_conv2, zip\_eq\_Nil\_iff |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
284 |
%\postw |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
285 |
% |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
286 |
%Sometimes Sledgehammer will helpfully suggest a missing assumption: |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
287 |
% |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
288 |
%\prew |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
289 |
%\slshape |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
290 |
%e: Candidate missing assumption: \\ |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
291 |
%length ys = length xs |
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
292 |
%\postw |
68565 | 293 |
|
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
294 |
\section{Hints} |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
295 |
\label{hints} |
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
296 |
|
42884 | 297 |
This section presents a few hints that should help you get the most out of |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46640
diff
changeset
|
298 |
Sledgehammer. Frequently asked questions are answered in |
45380 | 299 |
\S\ref{frequently-asked-questions}. |
42884 | 300 |
|
46242 | 301 |
%\newcommand\point[1]{\medskip\par{\sl\bfseries#1}\par\nopagebreak} |
75021 | 302 |
\newcommand\point[1]{\subsection{\slshape #1}} |
42763 | 303 |
|
68565 | 304 |
|
42763 | 305 |
\point{Presimplify the goal} |
306 |
||
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
307 |
For best results, first simplify your problem by calling \textit{auto} or at |
42945 | 308 |
least \textit{safe} followed by \textit{simp\_all}. The SMT solvers provide |
309 |
arithmetic decision procedures, but the ATPs typically do not (or if they do, |
|
310 |
Sledgehammer does not use it yet). Apart from Waldmeister, they are not |
|
53759
a198ce71de11
took out Waldmeister from list of default provers -- it's usually just visual noise, and its integration in Sledgehammer leaves much to be desired
blanchet
parents:
53757
diff
changeset
|
311 |
particularly good at heavy rewriting, but because they regard equations as |
42945 | 312 |
undirected, they often prove theorems that require the reverse orientation of a |
313 |
\textit{simp} rule. Higher-order problems can be tackled, but the success rate |
|
75021 | 314 |
is better for first-order problems. |
37517
19ba7ec5f1e3
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet
parents:
37498
diff
changeset
|
315 |
|
68565 | 316 |
|
53760
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
317 |
\point{Familiarize yourself with the main options} |
42763 | 318 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
319 |
Sledgehammer's options are fully documented in \S\ref{option-reference}. Many of |
42763 | 320 |
the options are very specialized, but serious users of the tool should at least |
321 |
familiarize themselves with the following options: |
|
322 |
||
323 |
\begin{enum} |
|
45516 | 324 |
\item[\labelitemi] \textbf{\textit{provers}} (\S\ref{mode-of-operation}) specifies |
42884 | 325 |
the automatic provers (ATPs and SMT solvers) that should be run whenever |
75021 | 326 |
Sledgehammer is invoked (e.g., ``\textit{provers}~= \textit{cvc4 e |
327 |
vampire zipperposition\/}''). |
|
42763 | 328 |
|
48294 | 329 |
\item[\labelitemi] \textbf{\textit{max\_facts}} (\S\ref{relevance-filter}) |
42884 | 330 |
specifies the maximum number of facts that should be passed to the provers. By |
48294 | 331 |
default, the value is prover-dependent but varies between about 50 and 1000. If |
332 |
the provers time out, you can try lowering this value to, say, 25 or 50 and see |
|
42884 | 333 |
if that helps. |
42763 | 334 |
|
49919 | 335 |
\item[\labelitemi] \textbf{\textit{isar\_proofs}} (\S\ref{output-format}) specifies |
75021 | 336 |
that Isar proofs should be generated, in addition to one-line proofs. The length |
337 |
of the Isar proofs can be controlled by setting \textit{compress} |
|
338 |
(\S\ref{output-format}). |
|
43038 | 339 |
|
45516 | 340 |
\item[\labelitemi] \textbf{\textit{timeout}} (\S\ref{timeouts}) controls the |
61317 | 341 |
provers' time limit. It is set to 30 seconds by default. |
42763 | 342 |
\end{enum} |
343 |
||
42884 | 344 |
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
|
345 |
(\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
|
346 |
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
|
347 |
``$(\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
|
348 |
to ensure that certain facts are included, or simply ``$(\textit{my\_facts})$'' |
58090 | 349 |
to force Sledgehammer to run only with $\textit{my\_facts}$ (and any facts |
350 |
chained into the goal). |
|
42763 | 351 |
|
68565 | 352 |
|
42763 | 353 |
\section{Frequently Asked Questions} |
354 |
\label{frequently-asked-questions} |
|
355 |
||
42945 | 356 |
This sections answers frequently (and infrequently) asked questions about |
48387 | 357 |
Sledgehammer. It is a good idea to skim over it now even if you do not have any |
75021 | 358 |
questions at this stage. |
42945 | 359 |
|
68565 | 360 |
|
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
361 |
\point{Which facts are passed to the automatic provers?} |
42883 | 362 |
|
75021 | 363 |
Sledgehammer heuristically selects a subset of lemmas from the currently loaded |
364 |
libraries. The component that performs this selection is called \emph{relevance |
|
365 |
filter} (\S\ref{relevance-filter}). |
|
48387 | 366 |
|
367 |
\begin{enum} |
|
368 |
\item[\labelitemi] |
|
72589 | 369 |
The traditional relevance filter, \emph{MePo} |
370 |
(\underline{Me}ng--\underline{Pau}lson), assigns a score to every available |
|
371 |
fact (lemma, theorem, definition, or axiom) based upon how many constants that |
|
75019
30a619de7973
use same default for FO and HO provers w.r.t. induction principles, based on evaluation -- this also simplifies the code
blanchet
parents:
75016
diff
changeset
|
372 |
fact shares with the goal. This process iterates to include facts |
72589 | 373 |
relevant to those just accepted. The constants are weighted to give unusual |
75019
30a619de7973
use same default for FO and HO provers w.r.t. induction principles, based on evaluation -- this also simplifies the code
blanchet
parents:
75016
diff
changeset
|
374 |
ones greater significance. MePo copes best when the goal contains some |
72589 | 375 |
unusual constants; if all the constants are common, it is unable to |
376 |
discriminate among the hundreds of facts that are picked up. The filter is also |
|
377 |
memoryless: It has no information about how many times a particular fact has |
|
378 |
been used in a proof, and it cannot learn. |
|
48387 | 379 |
|
380 |
\item[\labelitemi] |
|
57272
fd539459a112
enabled MaSh by default -- set 'MaSh' to 'none' in Isabelle Plugin Options to disable
blanchet
parents:
57245
diff
changeset
|
381 |
An alternative to MePo is \emph{MaSh} (\underline{Ma}chine Learner for |
fd539459a112
enabled MaSh by default -- set 'MaSh' to 'none' in Isabelle Plugin Options to disable
blanchet
parents:
57245
diff
changeset
|
382 |
\underline{S}ledge\underline{h}ammer). It applies machine learning to the |
fd539459a112
enabled MaSh by default -- set 'MaSh' to 'none' in Isabelle Plugin Options to disable
blanchet
parents:
57245
diff
changeset
|
383 |
problem of finding relevant facts. |
48387 | 384 |
|
61043 | 385 |
\item[\labelitemi] The \emph{MeSh} filter combines MePo and MaSh. This is |
386 |
the default. |
|
48387 | 387 |
\end{enum} |
388 |
||
42883 | 389 |
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
|
390 |
some provers get overwhelmed more easily than others. You can show the number of |
42883 | 391 |
facts given using the \textit{verbose} option (\S\ref{output-format}) and the |
392 |
actual facts using \textit{debug} (\S\ref{output-format}). |
|
393 |
||
394 |
Sledgehammer is good at finding short proofs combining a handful of existing |
|
395 |
lemmas. If you are looking for longer proofs, you must typically restrict the |
|
48294 | 396 |
number of facts, by setting the \textit{max\_facts} option |
43574 | 397 |
(\S\ref{relevance-filter}) to, say, 25 or 50. |
42883 | 398 |
|
42996 | 399 |
You can also influence which facts are actually selected in a number of ways. If |
400 |
you simply want to ensure that a fact is included, you can specify it using the |
|
75021 | 401 |
syntax ``$(\textit{add}{:}~\textit{my\_facts})$''. For example: |
42996 | 402 |
% |
403 |
\prew |
|
404 |
\textbf{sledgehammer} (\textit{add}: \textit{hd.simps} \textit{tl.simps}) |
|
405 |
\postw |
|
406 |
% |
|
407 |
The specified facts then replace the least relevant facts that would otherwise be |
|
408 |
included; the other selected facts remain the same. |
|
409 |
If you want to direct the selection in a particular direction, you can specify |
|
410 |
the facts via \textbf{using}: |
|
411 |
% |
|
412 |
\prew |
|
413 |
\textbf{using} \textit{hd.simps} \textit{tl.simps} \\ |
|
414 |
\textbf{sledgehammer} |
|
415 |
\postw |
|
416 |
% |
|
417 |
The facts are then more likely to be selected than otherwise, and if they are |
|
75021 | 418 |
selected at a given iteration of MePo, they also influence which facts are |
419 |
selected at subsequent iterations. |
|
42996 | 420 |
|
68565 | 421 |
|
46300 | 422 |
\point{Why does Metis fail to reconstruct the proof?} |
423 |
||
424 |
There are many reasons. If Metis runs seemingly forever, that is a sign that the |
|
57736
5f37ef22f9af
update documentation after removal of 'min' subcommand
blanchet
parents:
57733
diff
changeset
|
425 |
proof is too difficult for it. Metis's search is complete for first-order logic |
68565 | 426 |
with equality, so if the proof was found by a superposition-based ATP such as |
75021 | 427 |
E, SPASS, or Vampire, Metis should \emph{eventually} find it---in principle. |
46300 | 428 |
|
429 |
In some rare cases, \textit{metis} fails fairly quickly, and you get the error |
|
57736
5f37ef22f9af
update documentation after removal of 'min' subcommand
blanchet
parents:
57733
diff
changeset
|
430 |
message ``One-line proof reconstruction failed.'' This indicates that |
5f37ef22f9af
update documentation after removal of 'min' subcommand
blanchet
parents:
57733
diff
changeset
|
431 |
Sledgehammer determined that the goal is provable, but the proof is, for |
5f37ef22f9af
update documentation after removal of 'min' subcommand
blanchet
parents:
57733
diff
changeset
|
432 |
technical reasons, beyond \textit{metis}'s power. You can then try again with |
5f37ef22f9af
update documentation after removal of 'min' subcommand
blanchet
parents:
57733
diff
changeset
|
433 |
the \textit{strict} option (\S\ref{problem-encoding}). |
46300 | 434 |
|
68565 | 435 |
|
72589 | 436 |
\point{What are the \textit{full\_types}, \textit{no\_types}, and \\ |
46298 | 437 |
\textit{mono\_tags} arguments to Metis?} |
42883 | 438 |
|
46298 | 439 |
The \textit{metis}~(\textit{full\_types}) proof method |
440 |
and its cousin \textit{metis}~(\textit{mono\_tags}) are fully-typed |
|
53760
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
441 |
versions of Metis. It is somewhat slower than \textit{metis}, but the proof |
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43217
diff
changeset
|
442 |
search is fully typed, and it also includes more powerful rules such as the |
45516 | 443 |
axiom ``$x = \const{True} \mathrel{\lor} x = \const{False}$'' for reasoning in |
75021 | 444 |
higher-order positions (e.g., in set comprehensions). The method is tried as a |
72592 | 445 |
fallback when \textit{metis} fails, and it is sometimes generated by |
446 |
Sledgehammer instead of \textit{metis} if the proof obviously requires type |
|
447 |
information or if \textit{metis} failed when Sledgehammer preplayed the proof. |
|
46298 | 448 |
% |
43229 | 449 |
At the other end of the soundness spectrum, \textit{metis} (\textit{no\_types}) |
450 |
uses no type information at all during the proof search, which is more efficient |
|
451 |
but often fails. Calls to \textit{metis} (\textit{no\_types}) are occasionally |
|
452 |
generated by Sledgehammer. |
|
46298 | 453 |
% |
454 |
See the \textit{type\_enc} option (\S\ref{problem-encoding}) for details. |
|
43229 | 455 |
|
46298 | 456 |
Incidentally, if you ever see warnings such as |
42883 | 457 |
|
458 |
\prew |
|
43007 | 459 |
\slshape |
63729 | 460 |
Metis: Falling back on ``\textit{metis} (\textit{full\_types})'' |
42883 | 461 |
\postw |
462 |
||
45380 | 463 |
for a successful \textit{metis} proof, you can advantageously pass the |
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43217
diff
changeset
|
464 |
\textit{full\_types} option to \textit{metis} directly. |
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43217
diff
changeset
|
465 |
|
68565 | 466 |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
73859
diff
changeset
|
467 |
\point{And what are the \textit{lifting} and \textit{opaque\_lifting} \\ |
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
73859
diff
changeset
|
468 |
arguments to Metis?} |
46298 | 469 |
|
470 |
Orthogonally to the encoding of types, it is important to choose an appropriate |
|
72592 | 471 |
translation of $\lambda$-abstractions. Metis supports three translation |
472 |
schemes, in decreasing order of power: Curry combinators (the default), |
|
46298 | 473 |
$\lambda$-lifting, and a ``hiding'' scheme that disables all reasoning under |
75021 | 474 |
$\lambda$-abstractions. See the \textit{lam\_trans} option |
72592 | 475 |
(\S\ref{problem-encoding}) for details. |
46298 | 476 |
|
68565 | 477 |
|
478 |
\point{Are the generated proofs minimal?} |
|
43036 | 479 |
|
43054 | 480 |
Automatic provers frequently use many more facts than are necessary. |
72592 | 481 |
Sledgehammer includes a proof minimization tool that takes a set of facts returned by |
57722 | 482 |
a given prover and repeatedly calls a prover or proof method with subsets of |
483 |
those facts to find a minimal set. Reducing the number of facts typically helps |
|
75021 | 484 |
reconstruction and declutters the proof documents. |
43036 | 485 |
|
68565 | 486 |
|
43008
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
blanchet
parents:
43007
diff
changeset
|
487 |
\point{A strange error occurred---what should I do?} |
42763 | 488 |
|
489 |
Sledgehammer tries to give informative error messages. Please report any strange |
|
63729 | 490 |
error to the author at \authoremail. |
42763 | 491 |
|
68565 | 492 |
|
42763 | 493 |
\point{Auto can solve it---why not Sledgehammer?} |
494 |
||
495 |
Problems can be easy for \textit{auto} and difficult for automatic provers, but |
|
48387 | 496 |
the reverse is also true, so do not be discouraged if your first attempts fail. |
39320 | 497 |
Because the system refers to all theorems known to Isabelle, it is particularly |
57040 | 498 |
suitable when your goal has a short proof but requires lemmas that you do not |
499 |
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
|
500 |
|
68565 | 501 |
|
42883 | 502 |
\point{Why are there so many options?} |
503 |
||
72589 | 504 |
Sledgehammer's philosophy is that it should work out of the box, without user |
505 |
guidance. Most of the options are meant to be used by the Sledgehammer |
|
506 |
developers for experiments. |
|
42883 | 507 |
|
68565 | 508 |
|
36926 | 509 |
\section{Command Syntax} |
510 |
\label{command-syntax} |
|
511 |
||
46242 | 512 |
\subsection{Sledgehammer} |
57040 | 513 |
\label{sledgehammer} |
46242 | 514 |
|
36926 | 515 |
Sledgehammer can be invoked at any point when there is an open goal by entering |
516 |
the \textbf{sledgehammer} command in the theory file. Its general syntax is as |
|
517 |
follows: |
|
518 |
||
519 |
\prew |
|
43216 | 520 |
\textbf{sledgehammer} \qty{subcommand}$^?$ \qty{options}$^?$ \qty{facts\_override}$^?$ \qty{num}$^?$ |
36926 | 521 |
\postw |
522 |
||
43216 | 523 |
In the general syntax, the \qty{subcommand} may be any of the following: |
36926 | 524 |
|
525 |
\begin{enum} |
|
45516 | 526 |
\item[\labelitemi] \textbf{\textit{run} (the default):} Runs Sledgehammer on |
43216 | 527 |
subgoal number \qty{num} (1 by default), with the given options and facts. |
36926 | 528 |
|
45516 | 529 |
\item[\labelitemi] \textbf{\textit{supported\_provers}:} Prints the list of |
41724 | 530 |
automatic provers supported by Sledgehammer. See \S\ref{installation} and |
531 |
\S\ref{mode-of-operation} for more information on how to install automatic |
|
532 |
provers. |
|
36926 | 533 |
|
48393 | 534 |
\item[\labelitemi] \textbf{\textit{refresh\_tptp}:} Refreshes the list of remote |
535 |
ATPs available at System\-On\-TPTP \cite{sutcliffe-2000}. |
|
536 |
\end{enum} |
|
537 |
||
49365 | 538 |
In addition, the following subcommands provide finer control over machine |
48393 | 539 |
learning with MaSh: |
540 |
||
541 |
\begin{enum} |
|
542 |
\item[\labelitemi] \textbf{\textit{unlearn}:} Resets MaSh, erasing any |
|
543 |
persistent state. |
|
48387 | 544 |
|
48393 | 545 |
\item[\labelitemi] \textbf{\textit{learn\_isar}:} Invokes MaSh on the current |
546 |
theory to process all the available facts, learning from their Isabelle/Isar |
|
547 |
proofs. This happens automatically at Sledgehammer invocations if the |
|
548 |
\textit{learn} option (\S\ref{relevance-filter}) is enabled. |
|
48387 | 549 |
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
550 |
\item[\labelitemi] \textbf{\textit{learn\_prover}:} Invokes MaSh on the current |
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
551 |
theory to process all the available facts, learning from proofs generated by |
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
552 |
automatic provers. The prover to use and its timeout can be set using the |
48393 | 553 |
\textit{prover} (\S\ref{mode-of-operation}) and \textit{timeout} |
66735 | 554 |
(\S\ref{timeouts}) options. It is recommended to perform learning using a |
555 |
first-order ATP (such as E, SPASS, and Vampire) as opposed to a |
|
48393 | 556 |
higher-order ATP or an SMT solver. |
557 |
||
558 |
\item[\labelitemi] \textbf{\textit{relearn\_isar}:} Same as \textit{unlearn} |
|
559 |
followed by \textit{learn\_isar}. |
|
560 |
||
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
561 |
\item[\labelitemi] \textbf{\textit{relearn\_prover}:} Same as \textit{unlearn} |
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
562 |
followed by \textit{learn\_prover}. |
36926 | 563 |
\end{enum} |
564 |
||
43216 | 565 |
Sledgehammer's behavior can be influenced by various \qty{options}, which can be |
566 |
specified in brackets after the \textbf{sledgehammer} command. The |
|
567 |
\qty{options} are a list of key--value pairs of the form ``[$k_1 = v_1, |
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
568 |
\ldots, k_n = v_n$]''. For Boolean options, ``= \textit{true\/}'' is optional. |
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50459
diff
changeset
|
569 |
For example: |
36926 | 570 |
|
571 |
\prew |
|
49919 | 572 |
\textbf{sledgehammer} [\textit{isar\_proofs}, \,\textit{timeout} = 120] |
36926 | 573 |
\postw |
574 |
||
575 |
Default values can be set using \textbf{sledgehammer\_\allowbreak params}: |
|
576 |
||
577 |
\prew |
|
43216 | 578 |
\textbf{sledgehammer\_params} \qty{options} |
36926 | 579 |
\postw |
580 |
||
581 |
The supported options are described in \S\ref{option-reference}. |
|
582 |
||
43216 | 583 |
The \qty{facts\_override} argument lets you alter the set of facts that go |
584 |
through the relevance filter. It may be of the form ``(\qty{facts})'', where |
|
585 |
\qty{facts} is a space-separated list of Isabelle facts (theorems, local |
|
36926 | 586 |
assumptions, etc.), in which case the relevance filter is bypassed and the given |
43216 | 587 |
facts are used. It may also be of the form ``(\textit{add}:\ \qty{facts\/_{\mathrm{1}}})'', |
588 |
``(\textit{del}:\ \qty{facts\/_{\mathrm{2}}})'', or ``(\textit{add}:\ \qty{facts\/_{\mathrm{1}}}\ |
|
589 |
\textit{del}:\ \qty{facts\/_{\mathrm{2}}})'', where the relevance filter is instructed to |
|
590 |
proceed as usual except that it should consider \qty{facts\/_{\mathrm{1}}} |
|
591 |
highly-relevant and \qty{facts\/_{\mathrm{2}}} fully irrelevant. |
|
36926 | 592 |
|
53760
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
593 |
If you use Isabelle/jEdit, Sledgehammer also provides an automatic mode that can |
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
594 |
be enabled via the ``Auto Sledgehammer'' option under ``Plugins > Plugin Options |
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
595 |
> Isabelle > General.'' For automatic runs, only the first prover set using |
75022
e9e27d2e61a1
updated documentation of 'slice' (now 'slices') option
blanchet
parents:
75021
diff
changeset
|
596 |
\textit{provers} (\S\ref{mode-of-operation}) is considered, |
e9e27d2e61a1
updated documentation of 'slice' (now 'slices') option
blanchet
parents:
75021
diff
changeset
|
597 |
\textit{dont\_slice} (\S\ref{timeouts}) is set, fewer facts are |
54114 | 598 |
passed to the prover, \textit{fact\_filter} (\S\ref{relevance-filter}) is set to |
599 |
\textit{mepo}, \textit{strict} (\S\ref{problem-encoding}) is enabled, |
|
600 |
\textit{verbose} (\S\ref{output-format}) and \textit{debug} |
|
60306
6b7c64ab8bd2
made Auto Sledgehammer behave more like the real thing
blanchet
parents:
60185
diff
changeset
|
601 |
(\S\ref{output-format}) are disabled, and \textit{timeout} (\S\ref{timeouts}) is |
54114 | 602 |
superseded by the ``Auto Time Limit'' option in jEdit. Sledgehammer's output is |
603 |
also more concise. |
|
39320 | 604 |
|
68565 | 605 |
|
46242 | 606 |
\subsection{Metis} |
57040 | 607 |
\label{metis} |
46242 | 608 |
|
43216 | 609 |
The \textit{metis} proof method has the syntax |
610 |
||
611 |
\prew |
|
45518 | 612 |
\textbf{\textit{metis}}~(\qty{options})${}^?$~\qty{facts}${}^?$ |
43216 | 613 |
\postw |
614 |
||
45518 | 615 |
where \qty{facts} is a list of arbitrary facts and \qty{options} is a |
616 |
comma-separated list consisting of at most one $\lambda$ translation scheme |
|
617 |
specification with the same semantics as Sledgehammer's \textit{lam\_trans} |
|
618 |
option (\S\ref{problem-encoding}) and at most one type encoding specification |
|
619 |
with the same semantics as Sledgehammer's \textit{type\_enc} option |
|
620 |
(\S\ref{problem-encoding}). |
|
621 |
% |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
73859
diff
changeset
|
622 |
The supported $\lambda$ translation schemes are \textit{opaque\_lifting}, |
46366 | 623 |
\textit{lifting}, and \textit{combs} (the default). |
45518 | 624 |
% |
625 |
All the untyped type encodings listed in \S\ref{problem-encoding} are supported. |
|
626 |
For convenience, the following aliases are provided: |
|
627 |
\begin{enum} |
|
48393 | 628 |
\item[\labelitemi] \textbf{\textit{full\_types}:} Alias for \textit{poly\_guards\_query}. |
629 |
\item[\labelitemi] \textbf{\textit{partial\_types}:} Alias for \textit{poly\_args}. |
|
630 |
\item[\labelitemi] \textbf{\textit{no\_types}:} Alias for \textit{erased}. |
|
45518 | 631 |
\end{enum} |
43216 | 632 |
|
81286 | 633 |
The \textit{metis} method also supports the Isabelle option |
81757
4d15005da582
tuned documentation and order of instantiated facts
Lukas Bartl
parents:
81613
diff
changeset
|
634 |
[[\textit{metis\_instantiate}]], which tells \textit{metis} to infer and |
4d15005da582
tuned documentation and order of instantiated facts
Lukas Bartl
parents:
81613
diff
changeset
|
635 |
suggest instantiations of facts using \textbf{of} from a successful proof. |
68565 | 636 |
|
36926 | 637 |
\section{Option Reference} |
638 |
\label{option-reference} |
|
639 |
||
43014 | 640 |
\def\defl{\{} |
641 |
\def\defr{\}} |
|
642 |
||
36926 | 643 |
\def\flushitem#1{\item[]\noindent\kern-\leftmargin \textbf{#1}} |
47036 | 644 |
\def\optrueonly#1{\flushitem{\textit{#1} $\bigl[$= \textit{true}$\bigr]$\enskip}\nopagebreak\\[\parskip]} |
43014 | 645 |
\def\optrue#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\enskip \defl\textit{true}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
646 |
\def\opfalse#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\enskip \defl\textit{false}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
|
647 |
\def\opsmart#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{smart\_bool}$\bigr]$\enskip \defl\textit{smart}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
|
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46366
diff
changeset
|
648 |
\def\opsmartx#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{smart\_bool}$\bigr]$\enskip \defl\textit{smart}\defr\\\hbox{}\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
78149
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
649 |
\def\opsmartfalse#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{smart\_bool}$\bigr]$\enskip \defl\textit{false}\defr\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]} |
36926 | 650 |
\def\opnodefault#1#2{\flushitem{\textit{#1} = \qtybf{#2}} \nopagebreak\\[\parskip]} |
43014 | 651 |
\def\opnodefaultbrk#1#2{\flushitem{$\bigl[$\textit{#1} =$\bigr]$ \qtybf{#2}} \nopagebreak\\[\parskip]} |
652 |
\def\opdefault#1#2#3{\flushitem{\textit{#1} = \qtybf{#2}\enskip \defl\textit{#3}\defr} \nopagebreak\\[\parskip]} |
|
36926 | 653 |
\def\oparg#1#2#3{\flushitem{\textit{#1} \qtybf{#2} = \qtybf{#3}} \nopagebreak\\[\parskip]} |
654 |
\def\opargbool#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]} |
|
43014 | 655 |
\def\opargboolorsmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{smart\_bool}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]} |
36926 | 656 |
|
657 |
Sledgehammer's options are categorized as follows:\ mode of operation |
|
38984 | 658 |
(\S\ref{mode-of-operation}), problem encoding (\S\ref{problem-encoding}), |
659 |
relevance filter (\S\ref{relevance-filter}), output format |
|
57241 | 660 |
(\S\ref{output-format}), regression testing (\S\ref{regression-testing}), |
661 |
and timeouts (\S\ref{timeouts}). |
|
36926 | 662 |
|
663 |
The descriptions below refer to the following syntactic quantities: |
|
664 |
||
665 |
\begin{enum} |
|
45516 | 666 |
\item[\labelitemi] \qtybf{string}: A string. |
667 |
\item[\labelitemi] \qtybf{bool\/}: \textit{true} or \textit{false}. |
|
668 |
\item[\labelitemi] \qtybf{smart\_bool\/}: \textit{true}, \textit{false}, or |
|
40203 | 669 |
\textit{smart}. |
45516 | 670 |
\item[\labelitemi] \qtybf{int\/}: An integer. |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54788
diff
changeset
|
671 |
\item[\labelitemi] \qtybf{float}: A floating-point number (e.g., 2.5 or 60) |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54788
diff
changeset
|
672 |
expressing a number of seconds. |
45516 | 673 |
\item[\labelitemi] \qtybf{float\_pair\/}: A pair of floating-point numbers |
40343
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
blanchet
parents:
40341
diff
changeset
|
674 |
(e.g., 0.6 0.95). |
45516 | 675 |
\item[\labelitemi] \qtybf{smart\_int\/}: An integer or \textit{smart}. |
36926 | 676 |
\end{enum} |
677 |
||
43217 | 678 |
Default values are indicated in curly brackets (\textrm{\{\}}). Boolean options |
61317 | 679 |
have a negative counterpart (e.g., \textit{minimize} vs.\ |
680 |
\textit{dont\_minimize}). When setting Boolean options or their negative |
|
47963 | 681 |
counterparts, ``= \textit{true\/}'' may be omitted. |
36926 | 682 |
|
68565 | 683 |
|
36926 | 684 |
\subsection{Mode of Operation} |
685 |
\label{mode-of-operation} |
|
686 |
||
687 |
\begin{enum} |
|
43014 | 688 |
\opnodefaultbrk{provers}{string} |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
689 |
Specifies the automatic provers to use as a space-separated list (e.g., |
68565 | 690 |
``\textit{cvc4}~\textit{e}~\textit{spass}~\textit{vampire\/}''). |
46299 | 691 |
Provers can be run locally or remotely; see \S\ref{installation} for |
75036 | 692 |
installation instructions. By default, \textit{provers} is set to the list of |
693 |
all installed local provers. |
|
46299 | 694 |
|
695 |
The following local provers are supported: |
|
36926 | 696 |
|
48701 | 697 |
\begin{sloppy} |
36926 | 698 |
\begin{enum} |
70937 | 699 |
\item[\labelitemi] \textbf{\textit{agsyhol}:} agsyHOL is an automatic |
70936
646651bcf261
updated nomenclature for TPTP languages to use modern three-symbol abbreviations (e.g. TF1)
blanchet
parents:
70935
diff
changeset
|
700 |
higher-order prover developed by Fredrik Lindblad \cite{agsyHOL}. To use |
70937 | 701 |
agsyHOL, set the environment variable \texttt{AGSYHOL\_HOME} to the directory |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
702 |
that contains the \texttt{agsyHOL} executable. |
52078 | 703 |
|
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46640
diff
changeset
|
704 |
\item[\labelitemi] \textbf{\textit{alt\_ergo}:} Alt-Ergo is a polymorphic |
52078 | 705 |
ATP developed by Bobot et al.\ \cite{alt-ergo}. |
70936
646651bcf261
updated nomenclature for TPTP languages to use modern three-symbol abbreviations (e.g. TF1)
blanchet
parents:
70935
diff
changeset
|
706 |
It supports the TPTP polymorphic typed first-order format (TF1) via Why3 |
53102 | 707 |
\cite{why3}. To use Alt-Ergo, set the environment variable \texttt{WHY3\_HOME} |
56379
d8ecce5d51cd
use Alt-Ergo 0.95.2, the latest and greatest version
blanchet
parents:
56378
diff
changeset
|
708 |
to the directory that contains the \texttt{why3} executable. Sledgehammer |
d8ecce5d51cd
use Alt-Ergo 0.95.2, the latest and greatest version
blanchet
parents:
56378
diff
changeset
|
709 |
requires Alt-Ergo 0.95.2 and Why3 0.83. |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46640
diff
changeset
|
710 |
|
74048 | 711 |
\item[\labelitemi] \textbf{\textit{cvc4}:} CVC4 is an SMT solver developed by |
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
712 |
Barrett et al.\ \cite{cvc4}. To use CVC4, |
74048 | 713 |
set the environment variable \texttt{CVC4\_SOLVER} to the complete path of the |
714 |
executable, including the file name, or install the prebuilt CVC4 package from |
|
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
715 |
\download. |
57241 | 716 |
|
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
717 |
\item[\labelitemi] \textbf{\textit{cvc5}:} cvc5 is an SMT solver developed by |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
718 |
Barbosa et al.\ \cite{barbosa-et-al-cvc5}. To use cvc5, |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
719 |
set the environment variable \texttt{CVC5\_SOLVER} to the complete path of the |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
720 |
executable, including the file name. |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
721 |
|
45516 | 722 |
\item[\labelitemi] \textbf{\textit{e}:} E is a first-order resolution prover |
72592 | 723 |
developed by Stephan Schulz \cite{schulz-2019}. To use E, set the environment |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
724 |
variable \texttt{E\_HOME} to the directory that contains the \texttt{eproof} |
52757 | 725 |
executable and \texttt{E\_VERSION} to the version number (e.g., ``1.8''), or |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
726 |
install the prebuilt E package from \download. |
48652 | 727 |
|
48701 | 728 |
\item[\labelitemi] \textbf{\textit{iprover}:} iProver is a pure |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
729 |
instantiation-based prover developed by Konstantin Korovin \cite{korovin-2009}. |
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
730 |
To use iProver, set the environment variable \texttt{IPROVER\_HOME} to the |
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
731 |
directory that contains the \texttt{iproveropt} executable. iProver depends on |
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
732 |
Vampire to clausify problems, so make sure that Vampire is installed as well. |
48701 | 733 |
|
45516 | 734 |
\item[\labelitemi] \textbf{\textit{leo2}:} LEO-II is an automatic |
44098 | 735 |
higher-order prover developed by Christoph Benzm\"uller et al.\ \cite{leo2}, |
70936
646651bcf261
updated nomenclature for TPTP languages to use modern three-symbol abbreviations (e.g. TF1)
blanchet
parents:
70935
diff
changeset
|
736 |
with support for the TPTP typed higher-order syntax (TH0). To use LEO-II, set |
46242 | 737 |
the environment variable \texttt{LEO2\_HOME} to the directory that contains the |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
738 |
\texttt{leo} executable. |
44098 | 739 |
|
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66735
diff
changeset
|
740 |
\item[\labelitemi] \textbf{\textit{leo3}:} Leo-III is an automatic |
75021 | 741 |
higher-order prover developed by Alexander Steen, Christoph Benzm\"uller, |
742 |
et al.\ \cite{leo3}, |
|
70936
646651bcf261
updated nomenclature for TPTP languages to use modern three-symbol abbreviations (e.g. TF1)
blanchet
parents:
70935
diff
changeset
|
743 |
with support for the TPTP typed higher-order syntax (TH0). To use Leo-III, set |
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66735
diff
changeset
|
744 |
the environment variable \texttt{LEO3\_HOME} to the directory that contains the |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
745 |
\texttt{leo3} executable. |
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66735
diff
changeset
|
746 |
|
45516 | 747 |
\item[\labelitemi] \textbf{\textit{satallax}:} Satallax is an automatic |
44098 | 748 |
higher-order prover developed by Chad Brown et al.\ \cite{satallax}, with |
70936
646651bcf261
updated nomenclature for TPTP languages to use modern three-symbol abbreviations (e.g. TF1)
blanchet
parents:
70935
diff
changeset
|
749 |
support for the TPTP typed higher-order syntax (TH0). To use Satallax, set the |
46242 | 750 |
environment variable \texttt{SATALLAX\_HOME} to the directory that contains the |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
751 |
\texttt{satallax} executable. |
44098 | 752 |
|
45516 | 753 |
\item[\labelitemi] \textbf{\textit{spass}:} SPASS is a first-order resolution |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
754 |
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
|
755 |
To use SPASS, set the environment variable \texttt{SPASS\_HOME} to the directory |
47056 | 756 |
that contains the \texttt{SPASS} executable and \texttt{SPASS\_VERSION} to the |
47577 | 757 |
version number (e.g., ``3.8ds''), or install the prebuilt SPASS package from |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
758 |
\download. |
36926 | 759 |
|
48652 | 760 |
\item[\labelitemi] \textbf{\textit{vampire}:} Vampire is a first-order |
761 |
resolution prover developed by Andrei Voronkov and his colleagues |
|
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
762 |
\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
|
763 |
\texttt{VAMPIRE\_HOME} to the directory that contains the \texttt{vampire} |
48006 | 764 |
executable and \texttt{VAMPIRE\_VERSION} to the version number (e.g., |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
765 |
``4.2.2''). |
40942 | 766 |
|
59035
3a2153676705
renamed 'veriT' to 'verit', to stick to all-lowercase rule for prover names
blanchet
parents:
59034
diff
changeset
|
767 |
\item[\labelitemi] \textbf{\textit{verit}:} veriT \cite{bouton-et-al-2009} is an |
65516 | 768 |
SMT solver developed by David D\'eharbe, Pascal Fontaine, and their colleagues. |
72592 | 769 |
It is designed to produce detailed proofs for reconstruction in proof |
74388
d5e034f2c109
fixed veriT environment variable in sledgehammer's documentation
desharna
parents:
74367
diff
changeset
|
770 |
assistants. To use veriT, set the environment variable \texttt{ISABELLE\_VERIT} |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
771 |
to the complete path of the executable, including the file name. |
58497
20aaa307c0ff
updated docs with two provers: veriT and Zipperposition
blanchet
parents:
58090
diff
changeset
|
772 |
|
45516 | 773 |
\item[\labelitemi] \textbf{\textit{z3}:} Z3 is an SMT solver developed at |
72592 | 774 |
Microsoft Research \cite{de-moura-2008}. To use Z3, set the environment variable |
59961 | 775 |
\texttt{Z3\_SOLVER} to the complete path of the executable, including the |
74045
302994f5a3c2
updated Sledgehammer docs -- removed most version numbers since these are tedious to maintain
blanchet
parents:
74005
diff
changeset
|
776 |
file name. |
56378 | 777 |
|
58497
20aaa307c0ff
updated docs with two provers: veriT and Zipperposition
blanchet
parents:
58090
diff
changeset
|
778 |
\item[\labelitemi] \textbf{\textit{zipperposition}:} Zipperposition |
72174 | 779 |
\cite{cruanes-2014} is a higher-order superposition prover developed by Simon |
72589 | 780 |
Cruanes, Petar Vukmirovi\'c, and colleagues. To use Zipperposition, set the |
781 |
environment variable \texttt{ZIPPERPOSITION\_HOME} to the directory that |
|
782 |
contains the \texttt{zipperposition} executable and |
|
783 |
\texttt{ZIPPERPOSITION\_VERSION} to the version number (e.g., ``2.0.1''). |
|
57536 | 784 |
\end{enum} |
56378 | 785 |
|
48701 | 786 |
\end{sloppy} |
42945 | 787 |
|
57536 | 788 |
Moreover, the following remote provers are supported: |
42945 | 789 |
|
790 |
\begin{enum} |
|
52078 | 791 |
\item[\labelitemi] \textbf{\textit{remote\_agsyhol}:} The remote version of |
70937 | 792 |
agsyHOL runs on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}. |
793 |
||
794 |
\item[\labelitemi] \textbf{\textit{remote\_alt\_ergo}:} The remote version of |
|
795 |
Alt-Ergo runs on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}. |
|
52078 | 796 |
|
45516 | 797 |
\item[\labelitemi] \textbf{\textit{remote\_e}:} The remote version of E runs |
36926 | 798 |
on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}. |
799 |
||
48701 | 800 |
\item[\labelitemi] \textbf{\textit{remote\_iprover}:} The |
45339 | 801 |
remote version of iProver runs on Geoff Sutcliffe's Miami servers |
802 |
\cite{sutcliffe-2000}. |
|
803 |
||
45516 | 804 |
\item[\labelitemi] \textbf{\textit{remote\_leo2}:} The remote version of LEO-II |
44098 | 805 |
runs on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}. |
42964
bf45fd2488a2
document primitive support for LEO-II and Satallax
blanchet
parents:
42945
diff
changeset
|
806 |
|
67021
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66735
diff
changeset
|
807 |
\item[\labelitemi] \textbf{\textit{remote\_leo3}:} The remote version of Leo-III |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66735
diff
changeset
|
808 |
runs on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}. |
41f1f8c4259b
integrated Leo-III in Sledgehammer (thanks to Alexander Steen for the patch)
blanchet
parents:
66735
diff
changeset
|
809 |
|
45516 | 810 |
\item[\labelitemi] \textbf{\textit{remote\_waldmeister}:} Waldmeister is a unit |
42945 | 811 |
equality prover developed by Hillenbrand et al.\ \cite{waldmeister}. It can be |
43625 | 812 |
used to prove universally quantified equations using unconditional equations, |
813 |
corresponding to the TPTP CNF UEQ division. The remote version of Waldmeister |
|
814 |
runs on Geoff Sutcliffe's Miami servers. |
|
70940
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
70938
diff
changeset
|
815 |
|
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
70938
diff
changeset
|
816 |
\item[\labelitemi] \textbf{\textit{remote\_zipperposition}:} The remote |
ce3a05ad07b7
added support for Zipperposition on SystemOnTPTP
blanchet
parents:
70938
diff
changeset
|
817 |
version of Zipperposition runs on Geoff Sutcliffe's Miami servers. |
36926 | 818 |
\end{enum} |
819 |
||
75806
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
820 |
By default, Sledgehammer runs a subset of CVC4, E, SPASS, Vampire, veriT, Z3, |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
821 |
and Zipperposition in parallel, either locally or remotely---depending on the |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
822 |
number of processor cores available and on which provers are actually installed. |
2b106aae897c
added support for cvc5 (whose interface is almost identical to CVC4)
blanchet
parents:
75387
diff
changeset
|
823 |
It is generally beneficial to run several provers in parallel. |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
824 |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
825 |
\opnodefault{prover}{string} |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
826 |
Alias for \textit{provers}. |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39335
diff
changeset
|
827 |
|
81613
44afa6f1baad
added documentation for new Sledehammer option "cache_dir"
desharna
parents:
81612
diff
changeset
|
828 |
\opdefault{cache\_dir}{string}{""} |
44afa6f1baad
added documentation for new Sledehammer option "cache_dir"
desharna
parents:
81612
diff
changeset
|
829 |
Specifies whether Sledgehammer should cache the result of the external provers or not and, if yes, where. |
44afa6f1baad
added documentation for new Sledehammer option "cache_dir"
desharna
parents:
81612
diff
changeset
|
830 |
If the option is set to the empty string (i.e., ""), then no caching takes place. |
44afa6f1baad
added documentation for new Sledehammer option "cache_dir"
desharna
parents:
81612
diff
changeset
|
831 |
Otherwise, the string is interpreted as a path to a directory where the cached result will be saved. |
44afa6f1baad
added documentation for new Sledehammer option "cache_dir"
desharna
parents:
81612
diff
changeset
|
832 |
The content of the cache directory can be deleted at any time to reset the cache. |
44afa6f1baad
added documentation for new Sledehammer option "cache_dir"
desharna
parents:
81612
diff
changeset
|
833 |
|
78149
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
834 |
\opsmartfalse{falsify}{dont\_falsify} |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
835 |
Specifies whether Sledgehammer should look for falsifications or for proofs. If |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
836 |
the option is set to \textit{smart}, it looks for both. |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
837 |
|
77428 | 838 |
A falsification indicates that the goal, taken as an axiom, would be |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
839 |
inconsistent with some specific background facts if it were added as a lemma, |
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
840 |
indicating a likely issue with the goal. Sometimes the inconsistency involves |
77422 | 841 |
only the background theory; this might happen, for example, if a flawed axiom is |
842 |
used or if a flawed lemma was introduced with \textbf{sorry}. |
|
843 |
||
78149
d3122089b67c
disable 'falsify' and 'abduce' in Sledgehammer by default, since they don't seem to be very useful in practice
blanchet
parents:
77428
diff
changeset
|
844 |
\opdefault{abduce}{smart\_int}{0} |
77425 | 845 |
Specifies the maximum number of candidate missing assumptions that may be |
77422 | 846 |
displayed. These hypotheses are discovered heuristically by a process called |
77424 | 847 |
abduction (which stands in contrast to deduction)---that is, they are guessed |
848 |
and found to be sufficient to prove the goal. |
|
77422 | 849 |
|
850 |
Abduction is currently supported only by E. If the option is set to |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
851 |
\textit{smart}, abduction is enabled only in some of the E time slices, and at |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
852 |
most one candidate missing assumption is displayed. You can disable abduction |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
853 |
altogether by setting the option to 0 or enable it in all time slices by setting |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
854 |
it to a nonzero value. |
77269
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents:
75873
diff
changeset
|
855 |
|
77424 | 856 |
\optrueonly{dont\_abduce} |
857 |
Alias for ``\textit{abduce} = 0''. |
|
858 |
||
57722 | 859 |
\optrue{minimize}{dont\_minimize} |
72592 | 860 |
Specifies whether the proof minimization tool should be invoked automatically |
861 |
after proof search. |
|
45708
7c8bed80301f
updated Sledgehammer docs with new/renamed options
blanchet
parents:
45555
diff
changeset
|
862 |
|
7c8bed80301f
updated Sledgehammer docs with new/renamed options
blanchet
parents:
45555
diff
changeset
|
863 |
\nopagebreak |
47036 | 864 |
{\small See also \textit{preplay\_timeout} (\S\ref{timeouts}) |
865 |
and \textit{dont\_preplay} (\S\ref{timeouts}).} |
|
45708
7c8bed80301f
updated Sledgehammer docs with new/renamed options
blanchet
parents:
45555
diff
changeset
|
866 |
|
53801 | 867 |
\opfalse{spy}{dont\_spy} |
868 |
Specifies whether Sledgehammer should record statistics in |
|
869 |
\texttt{\$ISA\-BELLE\_\allowbreak HOME\_\allowbreak USER/\allowbreak spy\_\allowbreak sledgehammer}. |
|
870 |
These statistics can be useful to the developers of Sledgehammer. If you are willing to have your |
|
871 |
interactions recorded in the name of science, please enable this feature and send the statistics |
|
872 |
file every now and then to the author of this manual (\authoremail). |
|
873 |
To change the default value of this option globally, set the environment variable |
|
57107
2d502370ee99
changed MaSh to use SML version instead of Python version of naive Bayes by default (i.e. if MASH=yes in the settings, or 'fact_filter=mash' with no other explicit setting)
blanchet
parents:
57095
diff
changeset
|
874 |
\texttt{SLEDGEHAMMER\_SPY} to \textit{yes}. |
53801 | 875 |
|
876 |
\nopagebreak |
|
877 |
{\small See also \textit{debug} (\S\ref{output-format}).} |
|
878 |
||
36926 | 879 |
\opfalse{overlord}{no\_overlord} |
880 |
Specifies whether Sledgehammer should put its temporary files in |
|
881 |
\texttt{\$ISA\-BELLE\_\allowbreak HOME\_\allowbreak USER}, which is useful for |
|
882 |
debugging Sledgehammer but also unsafe if several instances of the tool are run |
|
48390
4147f2bc4442
add versioning to MaSh state + cleanup dead code
blanchet
parents:
48388
diff
changeset
|
883 |
simultaneously. The files are identified by the prefixes \texttt{prob\_} and |
4147f2bc4442
add versioning to MaSh state + cleanup dead code
blanchet
parents:
48388
diff
changeset
|
884 |
\texttt{mash\_}; you may safely remove them after Sledgehammer has run. |
36926 | 885 |
|
54139 | 886 |
\textbf{Warning:} This option is not thread-safe. Use at your own risks. |
887 |
||
36926 | 888 |
\nopagebreak |
889 |
{\small See also \textit{debug} (\S\ref{output-format}).} |
|
890 |
\end{enum} |
|
891 |
||
68565 | 892 |
|
48387 | 893 |
\subsection{Relevance Filter} |
894 |
\label{relevance-filter} |
|
895 |
||
896 |
\begin{enum} |
|
48388 | 897 |
\opdefault{fact\_filter}{string}{smart} |
898 |
Specifies the relevance filter to use. The following filters are available: |
|
899 |
||
900 |
\begin{enum} |
|
901 |
\item[\labelitemi] \textbf{\textit{mepo}:} |
|
902 |
The traditional memoryless MePo relevance filter. |
|
903 |
||
904 |
\item[\labelitemi] \textbf{\textit{mash}:} |
|
57532 | 905 |
The MaSh machine learner. Three learning algorithms are provided: |
57019 | 906 |
|
907 |
\begin{enum} |
|
57463 | 908 |
\item[\labelitemi] \textbf{\textit{nb}} is an implementation of naive Bayes. |
57019 | 909 |
|
57463 | 910 |
\item[\labelitemi] \textbf{\textit{knn}} is an implementation of $k$-nearest |
911 |
neighbors. |
|
912 |
||
57659
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
913 |
\item[\labelitemi] \textbf{\textit{nb\_knn}} (also called \textbf{\textit{yes}} |
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
914 |
and \textbf{\textit{sml}}) is a combination of naive Bayes and $k$-nearest |
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
915 |
neighbors. |
57019 | 916 |
\end{enum} |
917 |
||
57272
fd539459a112
enabled MaSh by default -- set 'MaSh' to 'none' in Isabelle Plugin Options to disable
blanchet
parents:
57245
diff
changeset
|
918 |
In addition, the special value \textit{none} is used to disable machine learning by |
fd539459a112
enabled MaSh by default -- set 'MaSh' to 'none' in Isabelle Plugin Options to disable
blanchet
parents:
57245
diff
changeset
|
919 |
default (cf.\ \textit{smart} below). |
fd539459a112
enabled MaSh by default -- set 'MaSh' to 'none' in Isabelle Plugin Options to disable
blanchet
parents:
57245
diff
changeset
|
920 |
|
57659
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
921 |
The default algorithm is \textit{nb\_knn}. The algorithm can be selected by |
61043 | 922 |
setting the ``MaSh'' option under ``Plugins > Plugin Options > Isabelle > |
57532 | 923 |
General'' in Isabelle/jEdit. Persistent data for both algorithms is stored in |
924 |
the directory \texttt{\$ISABELLE\_\allowbreak HOME\_\allowbreak USER/\allowbreak |
|
925 |
mash}. |
|
48388 | 926 |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
50929
diff
changeset
|
927 |
\item[\labelitemi] \textbf{\textit{mesh}:} The MeSh filter, which combines the |
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
50929
diff
changeset
|
928 |
rankings from MePo and MaSh. |
48388 | 929 |
|
57659
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
930 |
\item[\labelitemi] \textbf{\textit{smart}:} A combination of MePo, MaSh, and |
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
931 |
MeSh. If the learning algorithm is set to be \textit{none}, \textit{smart} |
b246943b3aa3
reenabled MaSh for Isabelle2014 release (hopefully)
blanchet
parents:
57636
diff
changeset
|
932 |
behaves like MePo. |
48388 | 933 |
\end{enum} |
934 |
||
48387 | 935 |
\opdefault{max\_facts}{smart\_int}{smart} |
936 |
Specifies the maximum number of facts that may be returned by the relevance |
|
55297
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
937 |
filter. If the option is set to \textit{smart} (the default), it effectively |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
938 |
takes a value that was empirically found to be appropriate for the prover. |
57107
2d502370ee99
changed MaSh to use SML version instead of Python version of naive Bayes by default (i.e. if MASH=yes in the settings, or 'fact_filter=mash' with no other explicit setting)
blanchet
parents:
57095
diff
changeset
|
939 |
Typical values lie between 50 and 1000. |
53757
8d1a059ebcdb
reduce the number of emitted MaSh commands (among others to facilitate debugging)
blanchet
parents:
53518
diff
changeset
|
940 |
|
48387 | 941 |
\opdefault{fact\_thresholds}{float\_pair}{\upshape 0.45~0.85} |
942 |
Specifies the thresholds above which facts are considered relevant by the |
|
943 |
relevance filter. The first threshold is used for the first iteration of the |
|
944 |
relevance filter and the second threshold is used for the last iteration (if it |
|
945 |
is reached). The effective threshold is quadratically interpolated for the other |
|
946 |
iterations. Each threshold ranges from 0 to 1, where 0 means that all theorems |
|
947 |
are relevant and 1 only theorems that refer to previously seen constants. |
|
948 |
||
48388 | 949 |
\optrue{learn}{dont\_learn} |
72592 | 950 |
Specifies whether Sledgehammer invocations should run MaSh to learn the |
53760
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
951 |
available theories (and hence provide more accurate results). Learning takes |
cf37f4b84824
moved focus to Isabell/jEdit and away from Proof General
blanchet
parents:
53759
diff
changeset
|
952 |
place only if MaSh is enabled. |
48388 | 953 |
|
48387 | 954 |
\opdefault{max\_new\_mono\_instances}{int}{smart} |
955 |
Specifies the maximum number of monomorphic instances to generate beyond |
|
956 |
\textit{max\_facts}. The higher this limit is, the more monomorphic instances |
|
957 |
are potentially generated. Whether monomorphization takes place depends on the |
|
55297
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
958 |
type encoding used. If the option is set to \textit{smart} (the default), it |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
959 |
takes a value that was empirically found to be appropriate for the prover. For |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
960 |
most provers, this value is 100. |
48387 | 961 |
|
962 |
\nopagebreak |
|
963 |
{\small See also \textit{type\_enc} (\S\ref{problem-encoding}).} |
|
964 |
||
965 |
\opdefault{max\_mono\_iters}{int}{smart} |
|
966 |
Specifies the maximum number of iterations for the monomorphization fixpoint |
|
967 |
construction. The higher this limit is, the more monomorphic instances are |
|
968 |
potentially generated. Whether monomorphization takes place depends on the |
|
55297
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
969 |
type encoding used. If the option is set to \textit{smart} (the default), it |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
970 |
takes a value that was empirically found to be appropriate for the prover. |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
971 |
For most provers, this value is 3. |
48387 | 972 |
|
973 |
\nopagebreak |
|
974 |
{\small See also \textit{type\_enc} (\S\ref{problem-encoding}).} |
|
73941
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
975 |
|
75019
30a619de7973
use same default for FO and HO provers w.r.t. induction principles, based on evaluation -- this also simplifies the code
blanchet
parents:
75016
diff
changeset
|
976 |
\opdefault{induction\_rules}{string}{exclude} |
73941
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
977 |
Specifies whether induction rules should be considered as relevant facts. |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
978 |
The following behaviors are available: |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
979 |
\begin{enum} |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
980 |
\item[\labelitemi] \textbf{\textit{exclude}:} |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
981 |
Induction rules are ignored by the relevance filter. |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
982 |
|
74957
089eeaaee525
proper documentation for induction_rules Sledgehammer option
desharna
parents:
74388
diff
changeset
|
983 |
\item[\labelitemi] \textbf{\textit{instantiate}:} |
75019
30a619de7973
use same default for FO and HO provers w.r.t. induction principles, based on evaluation -- this also simplifies the code
blanchet
parents:
75016
diff
changeset
|
984 |
Induction rules are instantiated based on the goal and then |
73941
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
985 |
considered by the relevance filter. |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
986 |
|
75019
30a619de7973
use same default for FO and HO provers w.r.t. induction principles, based on evaluation -- this also simplifies the code
blanchet
parents:
75016
diff
changeset
|
987 |
\item[\labelitemi] \textbf{\textit{include}:} |
30a619de7973
use same default for FO and HO provers w.r.t. induction principles, based on evaluation -- this also simplifies the code
blanchet
parents:
75016
diff
changeset
|
988 |
Induction rules are considered by the relevance filter. |
73941
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
989 |
\end{enum} |
48387 | 990 |
\end{enum} |
991 |
||
68565 | 992 |
|
36926 | 993 |
\subsection{Problem Encoding} |
994 |
\label{problem-encoding} |
|
995 |
||
45516 | 996 |
\newcommand\comb[1]{\const{#1}} |
997 |
||
36926 | 998 |
\begin{enum} |
45516 | 999 |
\opdefault{lam\_trans}{string}{smart} |
1000 |
Specifies the $\lambda$ translation scheme to use in ATP problems. The supported |
|
1001 |
translation schemes are listed below: |
|
1002 |
||
1003 |
\begin{enum} |
|
46366 | 1004 |
\item[\labelitemi] \textbf{\textit{lifting}:} Introduce a new |
45516 | 1005 |
supercombinator \const{c} for each cluster of $n$~$\lambda$-abstractions, |
1006 |
defined using an equation $\const{c}~x_1~\ldots~x_n = t$ ($\lambda$-lifting). |
|
1007 |
||
73935
269b2f976100
added documentation for changes to Sledgehammer option "lam_trans"
desharna
parents:
73932
diff
changeset
|
1008 |
\item[\labelitemi] \textbf{\textit{opaque\_lifting}:} Same as |
73941
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
1009 |
\textit{lifting}, except that the supercombinators are kept opaque, |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
1010 |
i.e. they are unspecified fresh constants. This effectively disables |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
1011 |
all reasoning under $\lambda$-abstractions. |
73935
269b2f976100
added documentation for changes to Sledgehammer option "lam_trans"
desharna
parents:
73932
diff
changeset
|
1012 |
|
46366 | 1013 |
\item[\labelitemi] \textbf{\textit{combs}:} Rewrite lambdas to the Curry |
45516 | 1014 |
combinators (\comb{I}, \comb{K}, \comb{S}, \comb{B}, \comb{C}). Combinators |
1015 |
enable the ATPs to synthesize $\lambda$-terms but tend to yield bulkier formulas |
|
1016 |
than $\lambda$-lifting: The translation is quadratic in the worst case, and the |
|
1017 |
equational definitions of the combinators are very prolific in the context of |
|
1018 |
resolution. |
|
1019 |
||
73935
269b2f976100
added documentation for changes to Sledgehammer option "lam_trans"
desharna
parents:
73932
diff
changeset
|
1020 |
\item[\labelitemi] \textbf{\textit{opaque\_combs}:} Same as |
73941
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
1021 |
\textit{combs}, except that the combinators are kept opaque, i.e. without |
bec00c7ef8dd
documented Sledgehammer option "induction_rules"
desharna
parents:
73935
diff
changeset
|
1022 |
equational definitions. |
73935
269b2f976100
added documentation for changes to Sledgehammer option "lam_trans"
desharna
parents:
73932
diff
changeset
|
1023 |
|
46366 | 1024 |
\item[\labelitemi] \textbf{\textit{combs\_and\_lifting}:} Introduce a new |
45516 | 1025 |
supercombinator \const{c} for each cluster of $\lambda$-abstractions and characterize it both using a |
1026 |
lifted equation $\const{c}~x_1~\ldots~x_n = t$ and via Curry combinators. |
|
1027 |
||
46366 | 1028 |
\item[\labelitemi] \textbf{\textit{combs\_or\_lifting}:} For each cluster of |
1029 |
$\lambda$-abstractions, heuristically choose between $\lambda$-lifting and Curry |
|
1030 |
combinators. |
|
1031 |
||
45516 | 1032 |
\item[\labelitemi] \textbf{\textit{keep\_lams}:} |
1033 |
Keep the $\lambda$-abstractions in the generated problems. This is available |
|
73859
bc263f1f68cd
added support for TFX's and THF's $ite to Sledgehammer
desharna
parents:
73858
diff
changeset
|
1034 |
only with provers that support $\lambda$s. |
45516 | 1035 |
|
1036 |
\item[\labelitemi] \textbf{\textit{smart}:} The actual translation scheme used |
|
1037 |
depends on the ATP and should be the most efficient scheme for that ATP. |
|
1038 |
\end{enum} |
|
1039 |
||
46366 | 1040 |
For SMT solvers, the $\lambda$ translation scheme is always \textit{lifting}, |
1041 |
irrespective of the value of this option. |
|
45516 | 1042 |
|
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46366
diff
changeset
|
1043 |
\opsmartx{uncurried\_aliases}{no\_uncurried\_aliases} |
46411 | 1044 |
Specifies whether fresh function symbols should be generated as aliases for |
1045 |
applications of curried functions in ATP problems. |
|
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46366
diff
changeset
|
1046 |
|
43627
ecd4bb7a8bc0
update documentation after "type_enc" renaming + fixed a few other out-of-date factlets
blanchet
parents:
43625
diff
changeset
|
1047 |
\opdefault{type\_enc}{string}{smart} |
ecd4bb7a8bc0
update documentation after "type_enc" renaming + fixed a few other out-of-date factlets
blanchet
parents:
43625
diff
changeset
|
1048 |
Specifies the type encoding to use in ATP problems. Some of the type encodings |
ecd4bb7a8bc0
update documentation after "type_enc" renaming + fixed a few other out-of-date factlets
blanchet
parents:
43625
diff
changeset
|
1049 |
are unsound, meaning that they can give rise to spurious proofs |
48093 | 1050 |
(unreconstructible using \textit{metis}). The type encodings are |
46300 | 1051 |
listed below, with an indication of their soundness in parentheses. |
48093 | 1052 |
An asterisk (*) indicates that the encoding is slightly incomplete for |
56120
04c37dfef684
updated Sledgehammer docs w.r.t. 'smt2' and 'z3_new'
blanchet
parents:
56119
diff
changeset
|
1053 |
reconstruction with \textit{metis}, unless the \textit{strict} option (described |
46302 | 1054 |
below) is enabled. |
42228 | 1055 |
|
1056 |
\begin{enum} |
|
48090 | 1057 |
\item[\labelitemi] \textbf{\textit{erased} (unsound):} No type information is |
46300 | 1058 |
supplied to the ATP, not even to resolve overloading. Types are simply erased. |
42582 | 1059 |
|
45516 | 1060 |
\item[\labelitemi] \textbf{\textit{poly\_guards} (sound):} Types are encoded using |
46300 | 1061 |
a predicate \const{g}$(\tau, t)$ that guards bound |
48090 | 1062 |
variables. Constants are annotated with their types, supplied as extra |
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
1063 |
arguments, to resolve overloading. |
42685 | 1064 |
|
45516 | 1065 |
\item[\labelitemi] \textbf{\textit{poly\_tags} (sound):} Each term and subterm is |
46300 | 1066 |
tagged with its type using a function $\const{t\/}(\tau, t)$. |
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
1067 |
|
45516 | 1068 |
\item[\labelitemi] \textbf{\textit{poly\_args} (unsound):} |
43990 | 1069 |
Like for \textit{poly\_guards} constants are annotated with their types to |
43002
e88fde86e4c2
mention contributions from LCP and explain metis and metisFT encodings
blanchet
parents:
42996
diff
changeset
|
1070 |
resolve overloading, but otherwise no type information is encoded. This |
57040 | 1071 |
is the default encoding used by the \textit{metis} proof method. |
42685 | 1072 |
|
45516 | 1073 |
\item[\labelitemi] |
42722 | 1074 |
\textbf{% |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1075 |
\textit{raw\_mono\_guards}, \textit{raw\_mono\_tags} (sound); \\ |
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1076 |
\textit{raw\_mono\_args} (unsound):} \\ |
43990 | 1077 |
Similar to \textit{poly\_guards}, \textit{poly\_tags}, and \textit{poly\_args}, |
42722 | 1078 |
respectively, but the problem is additionally monomorphized, meaning that type |
1079 |
variables are instantiated with heuristically chosen ground types. |
|
1080 |
Monomorphization can simplify reasoning but also leads to larger fact bases, |
|
1081 |
which can slow down the ATPs. |
|
42582 | 1082 |
|
45516 | 1083 |
\item[\labelitemi] |
42722 | 1084 |
\textbf{% |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1085 |
\textit{mono\_guards}, \textit{mono\_tags} (sound); |
72589 | 1086 |
\textit{mono\_args} \\ (unsound):} \\ |
42722 | 1087 |
Similar to |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1088 |
\textit{raw\_mono\_guards}, \textit{raw\_mono\_tags}, and |
72589 | 1089 |
\textit{raw\_mono\_\allowbreak args}, respectively but types are mangled in constant names |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1090 |
instead of being supplied as ground term arguments. The binary predicate |
46300 | 1091 |
$\const{g}(\tau, t)$ becomes a unary predicate |
1092 |
$\const{g\_}\tau(t)$, and the binary function |
|
1093 |
$\const{t}(\tau, t)$ becomes a unary function |
|
1094 |
$\const{t\_}\tau(t)$. |
|
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
|
1095 |
|
46435 | 1096 |
\item[\labelitemi] \textbf{\textit{mono\_native} (sound):} Exploits native |
70936
646651bcf261
updated nomenclature for TPTP languages to use modern three-symbol abbreviations (e.g. TF1)
blanchet
parents:
70935
diff
changeset
|
1097 |
first-order types if the prover supports the TF0, TF1, TH0, or TH1 syntax; |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46640
diff
changeset
|
1098 |
otherwise, falls back on \textit{mono\_guards}. The problem is monomorphized. |
43625 | 1099 |
|
72589 | 1100 |
\item[\labelitemi] \textbf{\textit{mono\_native\_fool} (sound):} Exploits native |
1101 |
first-order types, including Booleans, if the prover supports the TFX0, TFX1, |
|
1102 |
TH0, or TH1 syntax; otherwise, falls back on \textit{mono\_native}. The problem |
|
1103 |
is monomorphized. |
|
42681 | 1104 |
|
72589 | 1105 |
\item[\labelitemi] \textbf{\textit{mono\_native\_higher}, |
1106 |
\textit{mono\_native\_higher\_fool} \\ (sound):} Exploits native higher-order |
|
1107 |
types, including Booleans if ending with ``\textit{\_fool}'', if the prover |
|
1108 |
supports the TH0 syntax; otherwise, falls back on \textit{mono\_native} or |
|
1109 |
\textit{mono\_native\_fool}. The problem is monomorphized. |
|
1110 |
||
1111 |
\item[\labelitemi] \textbf{\textit{poly\_native}, \textit{poly\_native\_fool}, |
|
1112 |
\textit{poly\_native\_higher}, \\ \textit{poly\_native\_higher\_fool} (sound):} |
|
1113 |
Exploits native first-order polymorphic types if the prover supports the TF1, |
|
1114 |
TFX1, or TH1 syntax; otherwise, falls back on \textit{mono\_native}, |
|
1115 |
\textit{mono\_native\_fool}, \textit{mono\_native\_higher}, or |
|
1116 |
\textit{mono\_native\_higher\_fool}. |
|
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46640
diff
changeset
|
1117 |
|
45516 | 1118 |
\item[\labelitemi] |
42681 | 1119 |
\textbf{% |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1120 |
\textit{poly\_guards}?, \textit{poly\_tags}?, \textit{raw\_mono\_guards}?, \\ |
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1121 |
\textit{raw\_mono\_tags}?, \textit{mono\_guards}?, \textit{mono\_tags}?, \\ |
46435 | 1122 |
\textit{mono\_native}? (sound*):} \\ |
43990 | 1123 |
The type encodings \textit{poly\_guards}, \textit{poly\_tags}, |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44423
diff
changeset
|
1124 |
\textit{raw\_mono\_guards}, \textit{raw\_mono\_tags}, \textit{mono\_guards}, |
47036 | 1125 |
\textit{mono\_tags}, and \textit{mono\_native} are fully typed and sound. For |
1126 |
each of these, Sledgehammer also provides a lighter variant identified by a |
|
1127 |
question mark (`\hbox{?}')\ that detects and erases monotonic types, notably |
|
1128 |
infinite types. (For \textit{mono\_native}, the types are not actually erased |
|
1129 |
but rather replaced by a shared uniform type of individuals.) As argument to the |
|
1130 |
\textit{metis} proof method, the question mark is replaced by a |
|
1131 |
\hbox{``\textit{\_query\/}''} suffix. |
|
42582 | 1132 |
|
45516 | 1133 |
\item[\labelitemi] |
42887
771be1dcfef6
document new type system and soundness properties of the different systems
blanchet
parents:
42884
diff
changeset
|
1134 |
\textbf{% |
44769 | 1135 |
\textit{poly\_guards}??, \textit{poly\_tags}??, \textit{raw\_mono\_guards}??, \\ |
1136 |
\textit{raw\_mono\_tags}??, \textit{mono\_guards}??, \textit{mono\_tags}?? \\ |
|
46300 | 1137 |
(sound*):} \\ |
44816 | 1138 |
Even lighter versions of the `\hbox{?}' encodings. As argument to the |
1139 |
\textit{metis} proof method, the `\hbox{??}' suffix is replaced by |
|
46242 | 1140 |
\hbox{``\textit{\_query\_query\/}''}. |
44816 | 1141 |
|
45516 | 1142 |
\item[\labelitemi] |
44816 | 1143 |
\textbf{% |
48184 | 1144 |
\textit{poly\_guards}@, \textit{poly\_tags}@, \textit{raw\_mono\_guards}@, \\ |
1145 |
\textit{raw\_mono\_tags}@ (sound*):} \\ |
|
44816 | 1146 |
Alternative versions of the `\hbox{??}' encodings. As argument to the |
48184 | 1147 |
\textit{metis} proof method, the `\hbox{@}' suffix is replaced by |
1148 |
\hbox{``\textit{\_at\/}''}. |
|
44769 | 1149 |
|
48093 | 1150 |
\item[\labelitemi] \textbf{\textit{poly\_args}?, \textit{raw\_mono\_args}? (unsound):} \\ |
1151 |
Lighter versions of \textit{poly\_args} and \textit{raw\_mono\_args}. |
|
1152 |
||
45516 | 1153 |
\item[\labelitemi] \textbf{\textit{smart}:} The actual encoding used depends on |
47036 | 1154 |
the ATP and should be the most efficient sound encoding for that ATP. |
42228 | 1155 |
\end{enum} |
1156 |
||
46435 | 1157 |
For SMT solvers, the type encoding is always \textit{mono\_native}, irrespective |
44743 | 1158 |
of the value of this option. |
42888 | 1159 |
|
1160 |
\nopagebreak |
|
1161 |
{\small See also \textit{max\_new\_mono\_instances} (\S\ref{relevance-filter}) |
|
1162 |
and \textit{max\_mono\_iters} (\S\ref{relevance-filter}).} |
|
43574 | 1163 |
|
46302 | 1164 |
\opfalse{strict}{non\_strict} |
46300 | 1165 |
Specifies whether Sledgehammer should run in its strict mode. In that mode, |
46302 | 1166 |
sound type encodings marked with an asterisk (*) above are made complete |
46300 | 1167 |
for reconstruction with \textit{metis}, at the cost of some clutter in the |
1168 |
generated problems. This option has no effect if \textit{type\_enc} is |
|
1169 |
deliberately set to an unsound encoding. |
|
38591 | 1170 |
\end{enum} |
36926 | 1171 |
|
68565 | 1172 |
|
36926 | 1173 |
\subsection{Output Format} |
1174 |
\label{output-format} |
|
1175 |
||
1176 |
\begin{enum} |
|
1177 |
||
1178 |
\opfalse{verbose}{quiet} |
|
1179 |
Specifies whether the \textbf{sledgehammer} command should explain what it does. |
|
1180 |
||
1181 |
\opfalse{debug}{no\_debug} |
|
40203 | 1182 |
Specifies whether Sledgehammer should display additional debugging information |
1183 |
beyond what \textit{verbose} already displays. Enabling \textit{debug} also |
|
61317 | 1184 |
enables \textit{verbose} behind the scenes. |
36926 | 1185 |
|
1186 |
\nopagebreak |
|
53801 | 1187 |
{\small See also \textit{spy} (\S\ref{mode-of-operation}) and |
1188 |
\textit{overlord} (\S\ref{mode-of-operation}).} |
|
36926 | 1189 |
|
75030 | 1190 |
\opdefault{max\_proofs}{int}{\upshape 4} |
75031 | 1191 |
Specifies the maximum number of proofs to display before stopping. This is a |
1192 |
soft limit. |
|
75030 | 1193 |
|
51190
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents:
51189
diff
changeset
|
1194 |
\opsmart{isar\_proofs}{no\_isar\_proofs} |
55297
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
1195 |
Specifies whether Isar proofs should be output in addition to one-line proofs. |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
1196 |
The construction of Isar proof is still experimental and may sometimes fail; |
75873 | 1197 |
however, when they succeed they can be faster and sometimes more intelligible |
1198 |
than one-line proofs. If the option is set to \textit{smart} (the default), Isar |
|
1199 |
proofs are generated only when no working one-line proof is available. |
|
36926 | 1200 |
|
57784 | 1201 |
\opdefault{compress}{int}{smart} |
49919 | 1202 |
Specifies the granularity of the generated Isar proofs if \textit{isar\_proofs} |
51190
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents:
51189
diff
changeset
|
1203 |
is explicitly enabled. A value of $n$ indicates that each Isar proof step should |
57784 | 1204 |
correspond to a group of up to $n$ consecutive proof steps in the ATP proof. If |
1205 |
the option is set to \textit{smart} (the default), the compression factor is 10 |
|
1206 |
if the \textit{isar\_proofs} option is explicitly enabled; otherwise, it is |
|
1207 |
$\infty$. |
|
51189 | 1208 |
|
57245 | 1209 |
\optrueonly{dont\_compress} |
57784 | 1210 |
Alias for ``\textit{compress} = 1''. |
51189 | 1211 |
|
57245 | 1212 |
\optrue{try0}{dont\_try0} |
53765 | 1213 |
Specifies whether standard proof methods such as \textit{auto} and |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1214 |
\textit{blast} should be tried as alternatives to \textit{metis} or |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1215 |
\textit{smt}. The collection of methods is roughly the same as |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1216 |
for the \textbf{try0} command. |
55289 | 1217 |
|
71931
0c8a9c028304
simplified 'smt_proofs' option to be a binary option (instead of ternary), now that SMT proofs are accepted in the AFP (done with Martin Desharnais)
blanchet
parents:
70940
diff
changeset
|
1218 |
\optrue{smt\_proofs}{no\_smt\_proofs} |
61283 | 1219 |
Specifies whether the \textit{smt} proof method should be tried in addition to |
71931
0c8a9c028304
simplified 'smt_proofs' option to be a binary option (instead of ternary), now that SMT proofs are accepted in the AFP (done with Martin Desharnais)
blanchet
parents:
70940
diff
changeset
|
1220 |
Isabelle's built-in proof methods. |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1221 |
|
81757
4d15005da582
tuned documentation and order of instantiated facts
Lukas Bartl
parents:
81613
diff
changeset
|
1222 |
\opsmart{instantiate}{dont\_instantiate} |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1223 |
Specifies whether Metis should try to infer variable instantiations before proof |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1224 |
reconstruction, which results in instantiations of facts using \textbf{of} |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1225 |
(e.g. \textit{map\_prod\_surj\_on}[\textbf{of} \textit{f A} "\textit{f} |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1226 |
\textasciigrave \textit{A}" \textit{g B} "\textit{g} \textasciigrave |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1227 |
\textit{B}"]). This can make the proof methods faster and more intelligible. If |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1228 |
the option is set to \textit{smart} (the default), variable instantiations are |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80091
diff
changeset
|
1229 |
inferred only if proof reconstruction failed or timed out. |
36926 | 1230 |
\end{enum} |
1231 |
||
68565 | 1232 |
|
57241 | 1233 |
\subsection{Regression Testing} |
1234 |
\label{regression-testing} |
|
38984 | 1235 |
|
1236 |
\begin{enum} |
|
1237 |
\opnodefault{expect}{string} |
|
1238 |
Specifies the expected outcome, which must be one of the following: |
|
36926 | 1239 |
|
1240 |
\begin{enum} |
|
46300 | 1241 |
\item[\labelitemi] \textbf{\textit{some}:} Sledgehammer found a proof. |
75376 | 1242 |
\item[\labelitemi] \textbf{\textit{some\_preplayed}:} Sledgehammer found a proof that was successfully preplayed. |
45516 | 1243 |
\item[\labelitemi] \textbf{\textit{none}:} Sledgehammer found no proof. |
1244 |
\item[\labelitemi] \textbf{\textit{timeout}:} Sledgehammer timed out. |
|
77418
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
blanchet
parents:
77269
diff
changeset
|
1245 |
\item[\labelitemi] \textbf{\textit{resources\_out}:} Sledgehammer ran out of resources. |
45516 | 1246 |
\item[\labelitemi] \textbf{\textit{unknown}:} Sledgehammer encountered some |
40203 | 1247 |
problem. |
38984 | 1248 |
\end{enum} |
1249 |
||
61317 | 1250 |
Sledgehammer emits an error if the actual outcome differs from the expected outcome. This option is |
1251 |
useful for regression testing. |
|
38984 | 1252 |
|
75387 | 1253 |
The expected outcomes are not mutually exclusive. More specifically, \textit{some} is accepted |
1254 |
whenever \textit{some\_preplayed} is accepted as the former has strictly fewer requirements |
|
1255 |
than the later. |
|
1256 |
||
38984 | 1257 |
\nopagebreak |
61317 | 1258 |
{\small See also \textit{timeout} (\S\ref{timeouts}).} |
43038 | 1259 |
\end{enum} |
1260 |
||
68565 | 1261 |
|
43038 | 1262 |
\subsection{Timeouts} |
1263 |
\label{timeouts} |
|
1264 |
||
1265 |
\begin{enum} |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54788
diff
changeset
|
1266 |
\opdefault{timeout}{float}{\upshape 30} |
43038 | 1267 |
Specifies the maximum number of seconds that the automatic provers should spend |
1268 |
searching for a proof. This excludes problem preparation and is a soft limit. |
|
1269 |
||
75872
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1270 |
\opdefault{slices}{int}{\upshape 12 times the number of cores detected} |
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1271 |
Specifies the number of time slices. Time slices are the basic unit for prover |
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1272 |
invocations. They are divided among the available provers. A single prover |
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1273 |
invocation can occupy a single slice, two slices, or more, depending on the |
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1274 |
prover. Slicing (and thereby parallelism) can be disable by setting |
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1275 |
\textit{slices} to 1. Since slicing is a valuable optimization, you should |
8bfad7bc74cb
tweak Sledgehammer's slicing mechanism -- updated Zipperposition's slices and make them half as long as other provers' to pack more of them in 30 s
blanchet
parents:
75806
diff
changeset
|
1276 |
probably leave it enabled unless you are conducting experiments. |
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1277 |
|
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1278 |
\nopagebreak |
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1279 |
{\small See also \textit{verbose} (\S\ref{output-format}).} |
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1280 |
|
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1281 |
\optrueonly{dont\_slice} |
75022
e9e27d2e61a1
updated documentation of 'slice' (now 'slices') option
blanchet
parents:
75021
diff
changeset
|
1282 |
Alias for ``\textit{slices} = 1''. |
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1283 |
|
57719 | 1284 |
\opdefault{preplay\_timeout}{float}{\upshape 1} |
55297
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
1285 |
Specifies the maximum number of seconds that \textit{metis} or other proof |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
1286 |
methods should spend trying to ``preplay'' the found proof. If this option |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
1287 |
is set to 0, no preplaying takes place, and no timing information is displayed |
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents:
55290
diff
changeset
|
1288 |
next to the suggested proof method calls. |
45708
7c8bed80301f
updated Sledgehammer docs with new/renamed options
blanchet
parents:
45555
diff
changeset
|
1289 |
|
7c8bed80301f
updated Sledgehammer docs with new/renamed options
blanchet
parents:
45555
diff
changeset
|
1290 |
\nopagebreak |
7c8bed80301f
updated Sledgehammer docs with new/renamed options
blanchet
parents:
45555
diff
changeset
|
1291 |
{\small See also \textit{minimize} (\S\ref{mode-of-operation}).} |
47036 | 1292 |
|
1293 |
\optrueonly{dont\_preplay} |
|
1294 |
Alias for ``\textit{preplay\_timeout} = 0''. |
|
1295 |
||
36926 | 1296 |
\end{enum} |
1297 |
||
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1298 |
\section{Mirabelle Testing Tool} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1299 |
\label{mirabelle} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1300 |
|
72404 | 1301 |
The \texttt{isabelle mirabelle} tool executes Sledgehammer or other advisory |
73856 | 1302 |
tools (e.g., Nitpick) or tactics (e.g., \textit{auto}) on all subgoals emerging |
72404 | 1303 |
in a theory. It is typically used to measure the success rate of a proof tool |
1304 |
on some benchmark. Its command-line usage is as follows: |
|
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1305 |
|
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1306 |
{\small |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1307 |
\begin{verbatim} |
73855 | 1308 |
Usage: isabelle mirabelle [OPTIONS] [SESSIONS ...] |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1309 |
|
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1310 |
Options are: |
73855 | 1311 |
-A ACTION add to list of actions |
75021 | 1312 |
-O DIR output directory for log files (default: |
1313 |
"mirabelle") |
|
1314 |
-T THEORY theory restriction: NAME or |
|
1315 |
NAME[FIRST_LINE:LAST_LINE] |
|
1316 |
-m INT max. no. of calls to each action (0: unbounded) |
|
1317 |
(default 0) |
|
1318 |
-s INT run actions on every nth goal (0: uniform |
|
1319 |
distribution) (default 1) |
|
74077
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
desharna
parents:
74048
diff
changeset
|
1320 |
-t SECONDS timeout in seconds for each action (default 30) |
73855 | 1321 |
... |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1322 |
|
73855 | 1323 |
Apply the given ACTIONs at all theories and proof steps of the |
1324 |
specified sessions. |
|
1325 |
||
1326 |
The absence of theory restrictions (option -T) means to check all |
|
1327 |
theories fully. Otherwise only the named theories are checked. |
|
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1328 |
\end{verbatim} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1329 |
} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1330 |
|
73855 | 1331 |
Option \texttt{-A ACTION} specifies an action to run on all subgoals. When |
1332 |
specified multiple times, all actions are performed in parallel on all |
|
1333 |
selected subgoals. Available actions are \texttt{arith}, \texttt{metis}, |
|
1334 |
\texttt{quickcheck}, \texttt{sledgehammer}, \texttt{sledgehammer\_filter}, and |
|
1335 |
\texttt{try0}. |
|
1336 |
||
1337 |
Option \texttt{-O DIR} specifies the output directory, which is created |
|
1338 |
if needed. In this directory, a log file named "mirabelle.log" records the |
|
1339 |
position of each tested subgoal and the result of executing the actions. |
|
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1340 |
|
73855 | 1341 |
Option \texttt{-T THEORY} restricts the subgoals to those emerging from this |
1342 |
theory. When not provided, all subgoals from are theories are selected. When |
|
73856 | 1343 |
provided multiple times, the union of all specified theories' subgoals is |
73855 | 1344 |
selected. |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1345 |
|
73855 | 1346 |
Option \texttt{-m INT} specifies a maximum number of goals on which the action |
1347 |
are run. |
|
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1348 |
|
73855 | 1349 |
Option \texttt{-s INT} specifies a stride, effectively running the actions on |
75021 | 1350 |
every $n$th goal. |
73855 | 1351 |
|
1352 |
Option \texttt{-t SECONDS} specifies a generic timeout that the actions may |
|
72404 | 1353 |
interpret differently. |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1354 |
|
73855 | 1355 |
More specific documentation about low-level options, the \texttt{ACTION} |
1356 |
parameter, and its corresponding options can be found in the Isabelle tool |
|
72404 | 1357 |
usage by entering \texttt{isabelle mirabelle -?} on the command line. |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1358 |
|
73855 | 1359 |
The following subsections assume that the environment variable \texttt{AFP} is |
1360 |
defined and points to a release of the Archive of Formal Proofs. |
|
1361 |
||
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1362 |
|
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1363 |
\subsection{Example of Benchmarking Sledgehammer} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1364 |
|
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1365 |
\begin{verbatim} |
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1366 |
isabelle mirabelle -d '$AFP' -O output \ |
74958 | 1367 |
-A "sledgehammer[provers = e, timeout = 30]" \ |
73855 | 1368 |
VeriComp |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1369 |
\end{verbatim} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1370 |
|
73855 | 1371 |
This command specifies to run the Sledgehammer action, using the E prover with |
74958 | 1372 |
a timeout of 30 seconds, on all subgoals emerging from all theory in the AFP |
73855 | 1373 |
session VeriComp. The results are written to \texttt{output/mirabelle.log}. |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1374 |
|
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1375 |
\begin{verbatim} |
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1376 |
isabelle mirabelle -d '$AFP' -O output \ |
73855 | 1377 |
-T Semantics -T Compiler \ |
74958 | 1378 |
-A "sledgehammer[provers = e, timeout = 30]" \ |
73855 | 1379 |
VeriComp |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1380 |
\end{verbatim} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1381 |
|
73855 | 1382 |
This command also specifies to run the Sledgehammer action, but this time only |
1383 |
on subgoals emerging from theories Semantics or Compiler. |
|
1384 |
||
1385 |
||
1386 |
\subsection{Example of Benchmarking Multiple Tools} |
|
1387 |
||
1388 |
\begin{verbatim} |
|
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1389 |
isabelle mirabelle -d '$AFP' -O output -t 10 \ |
73857 | 1390 |
-A "try0" -A "metis" \ |
1391 |
VeriComp |
|
73855 | 1392 |
\end{verbatim} |
1393 |
||
1394 |
This command specifies two actions running the \textbf{try0} and \textbf{metis} |
|
73858 | 1395 |
commands, respectively, each with a timeout of 10 seconds. The results are |
73855 | 1396 |
written to \texttt{output/mirabelle.log}. |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1397 |
|
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1398 |
|
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1399 |
\subsection{Example of Generating TPTP Files} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1400 |
|
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1401 |
\begin{verbatim} |
75016
873b581fd690
generalized the 'slice' option towards more flexible slicing
blanchet
parents:
74981
diff
changeset
|
1402 |
isabelle mirabelle -d '$AFP' -O output \ |
74981
10df7a627ab6
split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents:
74958
diff
changeset
|
1403 |
-A "sledgehammer[provers = e, timeout = 5, keep_probs = true]" \ |
73855 | 1404 |
VeriComp |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1405 |
\end{verbatim} |
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1406 |
|
73855 | 1407 |
This command generates TPTP files using Sledgehammer. Since the file |
1408 |
is generated at the very beginning of every Sledgehammer invocation, |
|
1409 |
a timeout of five seconds making the prover fail faster speeds up |
|
74079
180ee02eb075
documented Mirabelle_Sledgehammer's new keep semantics
desharna
parents:
74077
diff
changeset
|
1410 |
processing the subgoals. The results are written in an action-specific |
180ee02eb075
documented Mirabelle_Sledgehammer's new keep semantics
desharna
parents:
74077
diff
changeset
|
1411 |
subdirectory of the specified output directory (\texttt{output}). A TPTP |
180ee02eb075
documented Mirabelle_Sledgehammer's new keep semantics
desharna
parents:
74077
diff
changeset
|
1412 |
file is generated for each subgoal. |
72342
4195e75a92ef
[mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents:
72319
diff
changeset
|
1413 |
|
36926 | 1414 |
\let\em=\sl |
48962
a1acc1cb0271
more standard document preparation within session context;
wenzelm
parents:
48803
diff
changeset
|
1415 |
\bibliography{manual}{} |
36926 | 1416 |
\bibliographystyle{abbrv} |
1417 |
||
1418 |
\end{document} |