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