doc-src/Ref/classical.tex
author paulson
Wed, 30 Apr 1997 16:36:59 +0200
changeset 3089 32dad29d4666
parent 2632 1612b99395d4
child 3108 335efc3f5632
permissions -rw-r--r--
Documented blast_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
     1
%% $Id$
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
     2
\chapter{The Classical Reasoner}\label{chap:classical}
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
     3
\index{classical reasoner|(}
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
     4
\newcommand\ainfer[2]{\begin{array}{r@{\,}l}#2\\ \hline#1\end{array}}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
     5
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
     6
Although Isabelle is generic, many users will be working in some extension
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
     7
of classical first-order logic.  Isabelle's set theory~{\tt ZF} is built
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
     8
upon theory~{\tt FOL}, while higher-order logic contains first-order logic
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
     9
as a fragment.  Theorem-proving in predicate logic is undecidable, but many
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    10
researchers have developed strategies to assist in this task.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    11
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
    12
Isabelle's classical reasoner is an \ML{} functor that accepts certain
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    13
information about a logic and delivers a suite of automatic tactics.  Each
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    14
tactic takes a collection of rules and executes a simple, non-clausal proof
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    15
procedure.  They are slow and simplistic compared with resolution theorem
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    16
provers, but they can save considerable time and effort.  They can prove
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    17
theorems such as Pelletier's~\cite{pelletier86} problems~40 and~41 in
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    18
seconds:
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    19
\[ (\exists y. \forall x. J(y,x) \bimp \neg J(x,x))  
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    20
   \imp  \neg (\forall x. \exists y. \forall z. J(z,y) \bimp \neg J(z,x)) \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    21
\[ (\forall z. \exists y. \forall x. F(x,y) \bimp F(x,z) \conj \neg F(x,x))
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    22
   \imp \neg (\exists z. \forall x. F(x,z))  
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    23
\]
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    24
%
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    25
The tactics are generic.  They are not restricted to first-order logic, and
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    26
have been heavily used in the development of Isabelle's set theory.  Few
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    27
interactive proof assistants provide this much automation.  The tactics can
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    28
be traced, and their components can be called directly; in this manner,
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    29
any proof can be viewed interactively.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    30
2479
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    31
The simplest way to apply the classical reasoner (to subgoal~$i$) is as
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    32
follows:
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    33
\begin{ttbox}
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
    34
by (Blast_tac \(i\));
2479
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    35
\end{ttbox}
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    36
If the subgoal is a simple formula of the predicate calculus or set theory,
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    37
then it should be proved quickly.  However, to use the classical reasoner
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
    38
effectively, you need to know how it works and how to choose among the many
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
    39
tactics available, including {\tt Fast_tac} and {\tt Best_tac}.
2479
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    40
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    41
We shall first discuss the underlying principles, then present the classical
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    42
reasoner.  Finally, we shall see how to instantiate it for new logics.  The
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
    43
logics {\tt FOL}, {\tt HOL} and {\tt ZF} have it already installed.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    44
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    45
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    46
\section{The sequent calculus}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    47
\index{sequent calculus}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    48
Isabelle supports natural deduction, which is easy to use for interactive
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    49
proof.  But natural deduction does not easily lend itself to automation,
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    50
and has a bias towards intuitionism.  For certain proofs in classical
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    51
logic, it can not be called natural.  The {\bf sequent calculus}, a
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    52
generalization of natural deduction, is easier to automate.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    53
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    54
A {\bf sequent} has the form $\Gamma\turn\Delta$, where $\Gamma$
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    55
and~$\Delta$ are sets of formulae.%
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    56
\footnote{For first-order logic, sequents can equivalently be made from
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    57
  lists or multisets of formulae.} The sequent
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    58
\[ P@1,\ldots,P@m\turn Q@1,\ldots,Q@n \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    59
is {\bf valid} if $P@1\conj\ldots\conj P@m$ implies $Q@1\disj\ldots\disj
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    60
Q@n$.  Thus $P@1,\ldots,P@m$ represent assumptions, each of which is true,
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    61
while $Q@1,\ldots,Q@n$ represent alternative goals.  A sequent is {\bf
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    62
basic} if its left and right sides have a common formula, as in $P,Q\turn
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    63
Q,R$; basic sequents are trivially valid.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    64
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    65
Sequent rules are classified as {\bf right} or {\bf left}, indicating which
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    66
side of the $\turn$~symbol they operate on.  Rules that operate on the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    67
right side are analogous to natural deduction's introduction rules, and
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    68
left rules are analogous to elimination rules.  
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    69
Recall the natural deduction rules for
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    70
  first-order logic, 
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    71
\iflabelundefined{fol-fig}{from {\it Introduction to Isabelle}}%
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    72
                          {Fig.\ts\ref{fol-fig}}.
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
    73
The sequent calculus analogue of~$({\imp}I)$ is the rule
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    74
$$ \ainfer{\Gamma &\turn \Delta, P\imp Q}{P,\Gamma &\turn \Delta,Q}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    75
   \eqno({\imp}R) $$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    76
This breaks down some implication on the right side of a sequent; $\Gamma$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    77
and $\Delta$ stand for the sets of formulae that are unaffected by the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    78
inference.  The analogue of the pair~$({\disj}I1)$ and~$({\disj}I2)$ is the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    79
single rule 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    80
$$ \ainfer{\Gamma &\turn \Delta, P\disj Q}{\Gamma &\turn \Delta,P,Q}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    81
   \eqno({\disj}R) $$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    82
This breaks down some disjunction on the right side, replacing it by both
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    83
disjuncts.  Thus, the sequent calculus is a kind of multiple-conclusion logic.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    84
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    85
To illustrate the use of multiple formulae on the right, let us prove
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    86
the classical theorem $(P\imp Q)\disj(Q\imp P)$.  Working backwards, we
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    87
reduce this formula to a basic sequent:
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    88
\[ \infer[(\disj)R]{\turn(P\imp Q)\disj(Q\imp P)}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    89
   {\infer[(\imp)R]{\turn(P\imp Q), (Q\imp P)\;}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    90
    {\infer[(\imp)R]{P \turn Q, (Q\imp P)\qquad}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    91
                    {P, Q \turn Q, P\qquad\qquad}}}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    92
\]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    93
This example is typical of the sequent calculus: start with the desired
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    94
theorem and apply rules backwards in a fairly arbitrary manner.  This yields a
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    95
surprisingly effective proof procedure.  Quantifiers add few complications,
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    96
since Isabelle handles parameters and schematic variables.  See Chapter~10
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    97
of {\em ML for the Working Programmer}~\cite{paulson91} for further
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    98
discussion.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
    99
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   100
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   101
\section{Simulating sequents by natural deduction}
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   102
Isabelle can represent sequents directly, as in the object-logic~{\tt LK}\@.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   103
But natural deduction is easier to work with, and most object-logics employ
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   104
it.  Fortunately, we can simulate the sequent $P@1,\ldots,P@m\turn
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   105
Q@1,\ldots,Q@n$ by the Isabelle formula
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   106
\[ \List{P@1;\ldots;P@m; \neg Q@2;\ldots; \neg Q@n}\Imp Q@1, \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   107
where the order of the assumptions and the choice of~$Q@1$ are arbitrary.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   108
Elim-resolution plays a key role in simulating sequent proofs.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   109
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   110
We can easily handle reasoning on the left.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   111
As discussed in
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   112
\iflabelundefined{destruct}{{\it Introduction to Isabelle}}{\S\ref{destruct}}, 
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   113
elim-resolution with the rules $(\disj E)$, $(\bot E)$ and $(\exists E)$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   114
achieves a similar effect as the corresponding sequent rules.  For the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   115
other connectives, we use sequent-style elimination rules instead of
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   116
destruction rules such as $({\conj}E1,2)$ and $(\forall E)$.  But note that
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   117
the rule $(\neg L)$ has no effect under our representation of sequents!
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   118
$$ \ainfer{\neg P,\Gamma &\turn \Delta}{\Gamma &\turn \Delta,P}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   119
   \eqno({\neg}L) $$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   120
What about reasoning on the right?  Introduction rules can only affect the
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   121
formula in the conclusion, namely~$Q@1$.  The other right-side formulae are
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   122
represented as negated assumptions, $\neg Q@2$, \ldots,~$\neg Q@n$.  
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   123
\index{assumptions!negated}
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   124
In order to operate on one of these, it must first be exchanged with~$Q@1$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   125
Elim-resolution with the {\bf swap} rule has this effect:
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   126
$$ \List{\neg P; \; \neg R\Imp P} \Imp R   \eqno(swap)$$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   127
To ensure that swaps occur only when necessary, each introduction rule is
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   128
converted into a swapped form: it is resolved with the second premise
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   129
of~$(swap)$.  The swapped form of~$({\conj}I)$, which might be
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   130
called~$({\neg\conj}E)$, is
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   131
\[ \List{\neg(P\conj Q); \; \neg R\Imp P; \; \neg R\Imp Q} \Imp R. \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   132
Similarly, the swapped form of~$({\imp}I)$ is
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   133
\[ \List{\neg(P\imp Q); \; \List{\neg R;P}\Imp Q} \Imp R  \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   134
Swapped introduction rules are applied using elim-resolution, which deletes
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   135
the negated formula.  Our representation of sequents also requires the use
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   136
of ordinary introduction rules.  If we had no regard for readability, we
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   137
could treat the right side more uniformly by representing sequents as
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   138
\[ \List{P@1;\ldots;P@m; \neg Q@1;\ldots; \neg Q@n}\Imp \bot. \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   139
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   140
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   141
\section{Extra rules for the sequent calculus}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   142
As mentioned, destruction rules such as $({\conj}E1,2)$ and $(\forall E)$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   143
must be replaced by sequent-style elimination rules.  In addition, we need
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   144
rules to embody the classical equivalence between $P\imp Q$ and $\neg P\disj
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   145
Q$.  The introduction rules~$({\disj}I1,2)$ are replaced by a rule that
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   146
simulates $({\disj}R)$:
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   147
\[ (\neg Q\Imp P) \Imp P\disj Q \]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   148
The destruction rule $({\imp}E)$ is replaced by
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   149
\[ \List{P\imp Q;\; \neg P\Imp R;\; Q\Imp R} \Imp R. \]
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   150
Quantifier replication also requires special rules.  In classical logic,
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   151
$\exists x{.}P$ is equivalent to $\neg\forall x{.}\neg P$; the rules
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   152
$(\exists R)$ and $(\forall L)$ are dual:
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   153
\[ \ainfer{\Gamma &\turn \Delta, \exists x{.}P}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   154
          {\Gamma &\turn \Delta, \exists x{.}P, P[t/x]} \; (\exists R)
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   155
   \qquad
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   156
   \ainfer{\forall x{.}P, \Gamma &\turn \Delta}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   157
          {P[t/x], \forall x{.}P, \Gamma &\turn \Delta} \; (\forall L)
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   158
\]
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   159
Thus both kinds of quantifier may be replicated.  Theorems requiring
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   160
multiple uses of a universal formula are easy to invent; consider 
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   161
\[ (\forall x.P(x)\imp P(f(x))) \conj P(a) \imp P(f^n(a)), \]
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   162
for any~$n>1$.  Natural examples of the multiple use of an existential
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   163
formula are rare; a standard one is $\exists x.\forall y. P(x)\imp P(y)$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   164
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   165
Forgoing quantifier replication loses completeness, but gains decidability,
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   166
since the search space becomes finite.  Many useful theorems can be proved
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   167
without replication, and the search generally delivers its verdict in a
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   168
reasonable time.  To adopt this approach, represent the sequent rules
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   169
$(\exists R)$, $(\exists L)$ and $(\forall R)$ by $(\exists I)$, $(\exists
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   170
E)$ and $(\forall I)$, respectively, and put $(\forall E)$ into elimination
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   171
form:
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   172
$$ \List{\forall x{.}P(x); P(t)\Imp Q} \Imp Q    \eqno(\forall E@2) $$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   173
Elim-resolution with this rule will delete the universal formula after a
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   174
single use.  To replicate universal quantifiers, replace the rule by
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   175
$$ \List{\forall x{.}P(x);\; \List{P(t); \forall x{.}P(x)}\Imp Q} \Imp Q.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   176
   \eqno(\forall E@3) $$
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   177
To replicate existential quantifiers, replace $(\exists I)$ by
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   178
\[ \List{\neg(\exists x{.}P(x)) \Imp P(t)} \Imp \exists x{.}P(x). \]
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   179
All introduction rules mentioned above are also useful in swapped form.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   180
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   181
Replication makes the search space infinite; we must apply the rules with
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   182
care.  The classical reasoner distinguishes between safe and unsafe
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   183
rules, applying the latter only when there is no alternative.  Depth-first
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   184
search may well go down a blind alley; best-first search is better behaved
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   185
in an infinite search space.  However, quantifier replication is too
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   186
expensive to prove any but the simplest theorems.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   187
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   188
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   189
\section{Classical rule sets}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   190
\index{classical sets}
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   191
Each automatic tactic takes a {\bf classical set} --- a collection of
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   192
rules, classified as introduction or elimination and as {\bf safe} or {\bf
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   193
unsafe}.  In general, safe rules can be attempted blindly, while unsafe
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   194
rules must be used with care.  A safe rule must never reduce a provable
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   195
goal to an unprovable set of subgoals.  
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   196
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   197
The rule~$({\disj}I1)$ is unsafe because it reduces $P\disj Q$ to~$P$.  Any
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   198
rule is unsafe whose premises contain new unknowns.  The elimination
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   199
rule~$(\forall E@2)$ is unsafe, since it is applied via elim-resolution,
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   200
which discards the assumption $\forall x{.}P(x)$ and replaces it by the
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   201
weaker assumption~$P(\Var{t})$.  The rule $({\exists}I)$ is unsafe for
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   202
similar reasons.  The rule~$(\forall E@3)$ is unsafe in a different sense:
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   203
since it keeps the assumption $\forall x{.}P(x)$, it is prone to looping.
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   204
In classical first-order logic, all rules are safe except those mentioned
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   205
above.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   206
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   207
The safe/unsafe distinction is vague, and may be regarded merely as a way
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   208
of giving some rules priority over others.  One could argue that
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   209
$({\disj}E)$ is unsafe, because repeated application of it could generate
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   210
exponentially many subgoals.  Induction rules are unsafe because inductive
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   211
proofs are difficult to set up automatically.  Any inference is unsafe that
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   212
instantiates an unknown in the proof state --- thus \ttindex{match_tac}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   213
must be used, rather than \ttindex{resolve_tac}.  Even proof by assumption
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   214
is unsafe if it instantiates unknowns shared with other subgoals --- thus
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   215
\ttindex{eq_assume_tac} must be used, rather than \ttindex{assume_tac}.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   216
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   217
\subsection{Adding rules to classical sets}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   218
Classical rule sets belong to the abstract type \mltydx{claset}, which
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   219
supports the following operations (provided the classical reasoner is
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   220
installed!):
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   221
\begin{ttbox} 
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   222
empty_cs    : claset
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   223
print_cs    : claset -> unit
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   224
addSIs      : claset * thm list -> claset                 \hfill{\bf infix 4}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   225
addSEs      : claset * thm list -> claset                 \hfill{\bf infix 4}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   226
addSDs      : claset * thm list -> claset                 \hfill{\bf infix 4}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   227
addIs       : claset * thm list -> claset                 \hfill{\bf infix 4}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   228
addEs       : claset * thm list -> claset                 \hfill{\bf infix 4}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   229
addDs       : claset * thm list -> claset                 \hfill{\bf infix 4}
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   230
delrules    : claset * thm list -> claset                 \hfill{\bf infix 4}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   231
\end{ttbox}
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   232
The add operations ignore any rule already present in the claset with the same
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   233
classification (such as Safe Introduction).  They print a warning if the rule
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   234
has already been added with some other classification, but add the rule
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   235
anyway.  Calling {\tt delrules} deletes all occurrences of a rule from the
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   236
claset, but see the warning below concerning destruction rules.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   237
\begin{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   238
\item[\ttindexbold{empty_cs}] is the empty classical set.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   239
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   240
\item[\ttindexbold{print_cs} $cs$] prints the rules of~$cs$.
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   241
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   242
\item[$cs$ addSIs $rules$] \indexbold{*addSIs}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   243
adds safe introduction~$rules$ to~$cs$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   244
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   245
\item[$cs$ addSEs $rules$] \indexbold{*addSEs}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   246
adds safe elimination~$rules$ to~$cs$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   247
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   248
\item[$cs$ addSDs $rules$] \indexbold{*addSDs}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   249
adds safe destruction~$rules$ to~$cs$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   250
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   251
\item[$cs$ addIs $rules$] \indexbold{*addIs}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   252
adds unsafe introduction~$rules$ to~$cs$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   253
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   254
\item[$cs$ addEs $rules$] \indexbold{*addEs}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   255
adds unsafe elimination~$rules$ to~$cs$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   256
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   257
\item[$cs$ addDs $rules$] \indexbold{*addDs}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   258
adds unsafe destruction~$rules$ to~$cs$.
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   259
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   260
\item[$cs$ delrules $rules$] \indexbold{*delrules}
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   261
deletes~$rules$ from~$cs$.  It prints a warning for those rules that are not
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   262
in~$cs$.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   263
\end{ttdescription}
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   264
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   265
\begin{warn}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   266
  If you added $rule$ using {\tt addSDs} or {\tt addDs}, then you must delete
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   267
  it as follows:
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   268
\begin{ttbox}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   269
\(cs\) delrules [make_elim \(rule\)]
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   270
\end{ttbox}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   271
\par\noindent
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   272
This is necessary because the operators {\tt addSDs} and {\tt addDs} convert
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   273
the destruction rules to elimination rules by applying \ttindex{make_elim},
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   274
and then insert them using {\tt addSEs} and {\tt addEs}, respectively.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   275
\end{warn}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   276
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   277
Introduction rules are those that can be applied using ordinary resolution.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   278
The classical set automatically generates their swapped forms, which will
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   279
be applied using elim-resolution.  Elimination rules are applied using
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   280
elim-resolution.  In a classical set, rules are sorted by the number of new
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   281
subgoals they will yield; rules that generate the fewest subgoals will be
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   282
tried first (see \S\ref{biresolve_tac}).
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   283
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   284
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   285
\subsection{Modifying the search step}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   286
For a given classical set, the proof strategy is simple.  Perform as many
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   287
safe inferences as possible; or else, apply certain safe rules, allowing
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   288
instantiation of unknowns; or else, apply an unsafe rule.  The tactics may
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   289
also apply {\tt hyp_subst_tac}, if they have been set up to do so (see
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   290
below).  They may perform a form of Modus Ponens: if there are assumptions
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   291
$P\imp Q$ and~$P$, then replace $P\imp Q$ by~$Q$.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   292
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   293
The classical reasoning tactics---except {\tt blast_tac}!---allow you to
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   294
modify this basic proof strategy by 
2632
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   295
applying two arbitrary {\bf wrapper tacticals} to it. This affects each step of
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   296
the search.  Usually they are the identity tacticals, but they could apply 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   297
another tactic before or after the step tactic. The first one, which is
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   298
considered to be safe, affects \ttindex{safe_step_tac} and all the tactics that
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   299
call it. The the second one, which may be unsafe, affects 
2632
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   300
\ttindex{step_tac}, \ttindex{slow_step_tac} and the tactics that call them.
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   301
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   302
\begin{ttbox} 
2632
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   303
addss        : claset * simpset -> claset                 \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   304
addSbefore   : claset *  (int -> tactic)  -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   305
addSaltern   : claset *  (int -> tactic)  -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   306
setSWrapper  : claset * ((int -> tactic) -> 
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   307
                         (int -> tactic)) -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   308
compSWrapper : claset * ((int -> tactic) -> 
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   309
                         (int -> tactic)) -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   310
addbefore    : claset *  (int -> tactic)  -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   311
addaltern    : claset *  (int -> tactic)  -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   312
setWrapper   : claset * ((int -> tactic) -> 
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   313
                         (int -> tactic)) -> claset       \hfill{\bf infix 4}
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   314
compWrapper  : claset * ((int -> tactic) -> 
1612b99395d4 corrected minor mistakes
oheimb
parents: 2631
diff changeset
   315
                         (int -> tactic)) -> claset       \hfill{\bf infix 4}
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   316
\end{ttbox}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   317
%
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   318
\index{simplification!from classical reasoner} 
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   319
The wrapper tacticals underly the operator \ttindex{addss}, which combines
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   320
each search step by simplification.  Strictly speaking, {\tt addss} is not
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   321
part of the classical reasoner.  It should be defined (using {\tt addSaltern (CHANGED o (safe_asm_more_full_simp_tac ss)}) when the simplifier is installed.
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   322
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   323
\begin{ttdescription}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   324
\item[$cs$ addss $ss$] \indexbold{*addss}
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   325
adds the simpset~$ss$ to the classical set. The assumptions and goal will be
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   326
simplified, in a safe way, after the safe steps of the search.
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   327
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   328
\item[$cs$ addSbefore $tac$] \indexbold{*addSbefore}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   329
changes the safe wrapper tactical to apply the given tactic {\em before}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   330
each safe step of the search.
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   331
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   332
\item[$cs$ addSaltern $tac$] \indexbold{*addSaltern}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   333
changes the safe wrapper tactical to apply the given tactic when a safe step 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   334
of the search would fail.
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   335
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   336
\item[$cs$ setSWrapper $tactical$] \indexbold{*setSWrapper}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   337
specifies a new safe wrapper tactical.  
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   338
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   339
\item[$cs$ compSWrapper $tactical$] \indexbold{*compSWrapper}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   340
composes the $tactical$ with the existing safe wrapper tactical, 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   341
to combine their effects. 
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   342
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   343
\item[$cs$ addbefore $tac$] \indexbold{*addbefore}
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   344
changes the (unsafe) wrapper tactical to apply the given tactic, which should
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   345
be safe, {\em before} each step of the search.
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   346
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   347
\item[$cs$ addaltern $tac$] \indexbold{*addaltern}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   348
changes the (unsafe) wrapper tactical to apply the given tactic 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   349
{\em alternatively} after each step of the search.
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   350
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   351
\item[$cs$ setWrapper $tactical$] \indexbold{*setWrapper}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   352
specifies a new (unsafe) wrapper tactical.  
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   353
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   354
\item[$cs$ compWrapper $tactical$] \indexbold{*compWrapper}
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   355
composes the $tactical$ with the existing (unsafe) wrapper tactical, 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   356
to combine their effects. 
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   357
\end{ttdescription}
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   358
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   359
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   360
\section{The classical tactics}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   361
\index{classical reasoner!tactics}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   362
If installed, the classical module provides several tactics (and other
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   363
operations) for simulating the classical sequent calculus.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   364
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   365
\subsection{The automatic tableau prover}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   366
The tactic {\tt blast_tac} finds a proof rapidly using a separate tableau
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   367
prover and then reconstructs the proof using Isabelle tactics.  It is much
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   368
faster and more powerful than the other classical reasoning tactics, but has
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   369
major limitations too.  
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   370
\begin{itemize}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   371
\item It does not use the wrapper tacticals described above, such as
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   372
  \ttindex{addss}.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   373
\item It ignores types, which can cause problems in \HOL.  If it applies a rule
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   374
  whose types are inappropriate, then proof reconstruction will fail.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   375
\item It does not perform higher-order unification, as needed by the rule {\tt
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   376
    rangeI} in {\HOL} and {\tt RepFunI} in {\ZF}.  There are often
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   377
    alternatives to such rules, for example {\tt
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   378
    range_eqI} and {\tt RepFun_eqI}.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   379
\item The message {\small\tt Function Var's argument not a bound variable\ }
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   380
relates to the lack of higher-order unification.  Function variables
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   381
may only be applied to parameters of the subgoal.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   382
\item Its proof strategy is more general than {\tt fast_tac}'s but can be
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   383
  slower.  If {\tt blast_tac} fails or seems to be running forever, try {\tt
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   384
  fast_tac} and the other tactics described below.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   385
\end{itemize}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   386
%
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   387
\begin{ttbox} 
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   388
blast_tac        : claset -> int -> tactic
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   389
Blast.depth_tac  : claset -> int -> int -> tactic
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   390
Blast.trace      : bool ref \hfill{\bf initially false}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   391
\end{ttbox}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   392
The two tactics differ on how they bound the number of unsafe steps used in a
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   393
proof.  While {\tt blast_tac} starts with a bound of zero and increases it
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   394
successively to~20, {\tt Blast.depth_tac} applies a user-supplied search bound.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   395
\begin{ttdescription}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   396
\item[\ttindexbold{blast_tac} $cs$ $i$] tries to prove
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   397
  subgoal~$i$ using iterative deepening to increase the search bound.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   398
  
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   399
\item[\ttindexbold{Blast.depth_tac} $cs$ $lim$ $i$] tries
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   400
  to prove subgoal~$i$ using a search bound of $lim$.  Often a slow
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   401
  proof using {\tt blast_tac} can be made much faster by supplying the
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   402
  successful search bound to this tactic instead.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   403
  
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   404
\item[\ttindexbold{Blast.trace} := true;] \index{tracing!of classical prover}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   405
  causes the tableau prover to print a trace of its search.  At each step it
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   406
  displays the formula currently being examined and reports whether the branch
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   407
  has been closed, extended or split.
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   408
\end{ttdescription}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   409
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   410
\subsection{The automatic tactics}
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   411
\begin{ttbox} 
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   412
fast_tac      : claset -> int -> tactic
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   413
best_tac      : claset -> int -> tactic
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   414
slow_tac      : claset -> int -> tactic
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   415
slow_best_tac : claset -> int -> tactic
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   416
\end{ttbox}
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   417
These tactics work by applying {\tt step_tac} or {\tt slow_step_tac}
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   418
repeatedly.  Their effect is restricted (by {\tt SELECT_GOAL}) to one subgoal;
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   419
they either prove this subgoal or fail.  The {\tt slow_} versions are more
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   420
powerful but can be much slower.  
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   421
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   422
The best-first tactics are guided by a heuristic function: typically, the
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   423
total size of the proof state.  This function is supplied in the functor call
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   424
that sets up the classical reasoner.
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   425
\begin{ttdescription}
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   426
\item[\ttindexbold{fast_tac} $cs$ $i$] applies {\tt step_tac} using
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   427
depth-first search, to prove subgoal~$i$.
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   428
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   429
\item[\ttindexbold{best_tac} $cs$ $i$] applies {\tt step_tac} using
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   430
best-first search, to prove subgoal~$i$.
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   431
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   432
\item[\ttindexbold{slow_tac} $cs$ $i$] applies {\tt slow_step_tac} using
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   433
depth-first search, to prove subgoal~$i$.
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   434
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   435
\item[\ttindexbold{slow_best_tac} $cs$ $i$] applies {\tt slow_step_tac} using
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   436
best-first search, to prove subgoal~$i$.
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   437
\end{ttdescription}
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   438
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   439
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   440
\subsection{Depth-limited tactics}
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   441
\begin{ttbox} 
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   442
depth_tac  : claset -> int -> int -> tactic
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   443
deepen_tac : claset -> int -> int -> tactic
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   444
\end{ttbox}
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   445
These work by exhaustive search up to a specified depth.  Unsafe rules are
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   446
modified to preserve the formula they act on, so that it be used repeatedly.
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   447
They can prove more goals than {\tt fast_tac} can but are much
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   448
slower, for example if the assumptions have many universal quantifiers.
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   449
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   450
The depth limits the number of unsafe steps.  If you can estimate the minimum
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   451
number of unsafe steps needed, supply this value as~$m$ to save time.
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   452
\begin{ttdescription}
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   453
\item[\ttindexbold{depth_tac} $cs$ $m$ $i$] 
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   454
tries to prove subgoal~$i$ by exhaustive search up to depth~$m$.
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   455
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   456
\item[\ttindexbold{deepen_tac} $cs$ $m$ $i$] 
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   457
tries to prove subgoal~$i$ by iterative deepening.  It calls {\tt depth_tac}
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   458
repeatedly with increasing depths, starting with~$m$.
332
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   459
\end{ttdescription}
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   460
01b87a921967 final Springer copy
lcp
parents: 319
diff changeset
   461
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   462
\subsection{Single-step tactics}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   463
\begin{ttbox} 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   464
safe_step_tac : claset -> int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   465
safe_tac      : claset        -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   466
inst_step_tac : claset -> int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   467
step_tac      : claset -> int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   468
slow_step_tac : claset -> int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   469
\end{ttbox}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   470
The automatic proof procedures call these tactics.  By calling them
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   471
yourself, you can execute these procedures one step at a time.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   472
\begin{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   473
\item[\ttindexbold{safe_step_tac} $cs$ $i$] performs a safe step on
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   474
subgoal~$i$. The safe wrapper tactical is applied to a tactic that may include 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   475
proof by assumption or Modus Ponens (taking care not to instantiate unknowns), 
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   476
or {\tt hyp_subst_tac}.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   477
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   478
\item[\ttindexbold{safe_tac} $cs$] repeatedly performs safe steps on all 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   479
subgoals.  It is deterministic, with at most one outcome.  If the automatic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   480
tactics fail, try using {\tt safe_tac} to open up your formula; then you
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   481
can replicate certain quantifiers explicitly by applying appropriate rules.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   482
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   483
\item[\ttindexbold{inst_step_tac} $cs$ $i$] is like {\tt safe_step_tac},
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   484
but allows unknowns to be instantiated.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   485
1099
f4335b56f772 Covers wrapper tacticals: setwrapper, ..., addss
lcp
parents: 875
diff changeset
   486
\item[\ttindexbold{step_tac} $cs$ $i$] is the basic step of the proof
2631
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   487
  procedure.  The (unsafe) wrapper tactical is applied to a tactic that tries
5e5c78ba955e description of safe vs. unsafe wrapper and the functions involved
oheimb
parents: 2479
diff changeset
   488
 {\tt safe_tac}, {\tt inst_step_tac}, or applies an unsafe rule from~$cs$.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   489
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   490
\item[\ttindexbold{slow_step_tac}] 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   491
  resembles {\tt step_tac}, but allows backtracking between using safe
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   492
  rules with instantiation ({\tt inst_step_tac}) and using unsafe rules.
875
a0b71a4bbe5e documented slow_tac, slow_best_tac, depth_tac, deepen_tac
lcp
parents: 332
diff changeset
   493
  The resulting search space is larger.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   494
\end{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   495
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   496
\subsection{The current claset}
2479
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
   497
Some logics (\FOL, {\HOL} and \ZF) support the concept of a current
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
   498
claset\index{claset!current}.  This is a default set of classical rules.  The
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
   499
underlying idea is quite similar to that of a current simpset described in
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
   500
\S\ref{sec:simp-for-dummies}; please read that section, including its
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
   501
warnings.  Just like simpsets, clasets can be associated with theories.  The
57109c1a653d Updated account of implicit simpsets and clasets
paulson
parents: 1869
diff changeset
   502
tactics
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   503
\begin{ttbox}
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   504
Blast_tac     : int -> tactic
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   505
Fast_tac     : int -> tactic
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   506
Best_tac     : int -> tactic
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   507
Deepen_tac   : int -> int -> tactic
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   508
Step_tac     : int -> tactic
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   509
\end{ttbox}
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   510
\indexbold{*Blast_tac}\indexbold{*Best_tac}\indexbold{*Fast_tac}%
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   511
\indexbold{*Deepen_tac}\indexbold{*Step_tac}
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   512
make use of the current claset. E.g. {\tt Blast_tac} is defined as follows:
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   513
\begin{ttbox}
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   514
fun Blast_tac i = fast_tac (!claset) i;
1869
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   515
\end{ttbox}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   516
where \ttindex{!claset} is the current claset.
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   517
The functions
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   518
\begin{ttbox}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   519
AddSIs, AddSEs, AddSDs, AddIs, AddEs, AddDs: thm list -> unit
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   520
\end{ttbox}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   521
\indexbold{*AddSIs} \indexbold{*AddSEs} \indexbold{*AddSDs}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   522
\indexbold{*AddIs} \indexbold{*AddEs} \indexbold{*AddDs}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   523
are used to add rules to the current claset. They work exactly like their
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   524
lower case counterparts {\tt addSIs} etc.
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   525
\begin{ttbox}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   526
Delrules : thm list -> unit
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   527
\end{ttbox}
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   528
deletes rules from the current claset. You do not need to worry via which of
065bd29adc7a Added section about current claset.
berghofe
parents: 1099
diff changeset
   529
the above {\tt Add} functions the rule was initially added.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   530
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   531
\subsection{Other useful tactics}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   532
\index{tactics!for contradiction}
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   533
\index{tactics!for Modus Ponens}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   534
\begin{ttbox} 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   535
contr_tac    :             int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   536
mp_tac       :             int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   537
eq_mp_tac    :             int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   538
swap_res_tac : thm list -> int -> tactic
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   539
\end{ttbox}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   540
These can be used in the body of a specialized search.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   541
\begin{ttdescription}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   542
\item[\ttindexbold{contr_tac} {\it i}]\index{assumptions!contradictory}
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   543
  solves subgoal~$i$ by detecting a contradiction among two assumptions of
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   544
  the form $P$ and~$\neg P$, or fail.  It may instantiate unknowns.  The
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   545
  tactic can produce multiple outcomes, enumerating all possible
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   546
  contradictions.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   547
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   548
\item[\ttindexbold{mp_tac} {\it i}] 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   549
is like {\tt contr_tac}, but also attempts to perform Modus Ponens in
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   550
subgoal~$i$.  If there are assumptions $P\imp Q$ and~$P$, then it replaces
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   551
$P\imp Q$ by~$Q$.  It may instantiate unknowns.  It fails if it can do
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   552
nothing.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   553
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   554
\item[\ttindexbold{eq_mp_tac} {\it i}] 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   555
is like {\tt mp_tac} {\it i}, but may not instantiate unknowns --- thus, it
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   556
is safe.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   557
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   558
\item[\ttindexbold{swap_res_tac} {\it thms} {\it i}] refines subgoal~$i$ of
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   559
the proof state using {\it thms}, which should be a list of introduction
3089
32dad29d4666 Documented blast_tac
paulson
parents: 2632
diff changeset
   560
rules.  First, it attempts to prove the goal using {\tt assume_tac} or
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   561
{\tt contr_tac}.  It then attempts to apply each rule in turn, attempting
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   562
resolution and also elim-resolution with the swapped form.
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   563
\end{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   564
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   565
\subsection{Creating swapped rules}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   566
\begin{ttbox} 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   567
swapify   : thm list -> thm list
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   568
joinrules : thm list * thm list -> (bool * thm) list
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   569
\end{ttbox}
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   570
\begin{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   571
\item[\ttindexbold{swapify} {\it thms}] returns a list consisting of the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   572
swapped versions of~{\it thms}, regarded as introduction rules.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   573
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   574
\item[\ttindexbold{joinrules} ({\it intrs}, {\it elims})]
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   575
joins introduction rules, their swapped versions, and elimination rules for
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   576
use with \ttindex{biresolve_tac}.  Each rule is paired with~{\tt false}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   577
(indicating ordinary resolution) or~{\tt true} (indicating
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   578
elim-resolution).
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   579
\end{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   580
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   581
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   582
\section{Setting up the classical reasoner}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   583
\index{classical reasoner!setting up}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   584
Isabelle's classical object-logics, including {\tt FOL} and {\tt HOL}, have
286
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   585
the classical reasoner already set up.  When defining a new classical logic,
e7efbf03562b first draft of Springer book
lcp
parents: 104
diff changeset
   586
you should set up the reasoner yourself.  It consists of the \ML{} functor
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   587
\ttindex{ClassicalFun}, which takes the argument
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   588
signature {\tt CLASSICAL_DATA}:
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   589
\begin{ttbox} 
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   590
signature CLASSICAL_DATA =
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   591
  sig
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   592
  val mp             : thm
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   593
  val not_elim       : thm
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   594
  val swap           : thm
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   595
  val sizef          : thm -> int
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   596
  val hyp_subst_tacs : (int -> tactic) list
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   597
  end;
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   598
\end{ttbox}
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   599
Thus, the functor requires the following items:
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   600
\begin{ttdescription}
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   601
\item[\tdxbold{mp}] should be the Modus Ponens rule
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   602
$\List{\Var{P}\imp\Var{Q};\; \Var{P}} \Imp \Var{Q}$.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   603
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   604
\item[\tdxbold{not_elim}] should be the contradiction rule
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   605
$\List{\neg\Var{P};\; \Var{P}} \Imp \Var{R}$.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   606
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   607
\item[\tdxbold{swap}] should be the swap rule
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   608
$\List{\neg \Var{P}; \; \neg \Var{R}\Imp \Var{P}} \Imp \Var{R}$.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   609
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   610
\item[\ttindexbold{sizef}] is the heuristic function used for best-first
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   611
search.  It should estimate the size of the remaining subgoals.  A good
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   612
heuristic function is \ttindex{size_of_thm}, which measures the size of the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   613
proof state.  Another size function might ignore certain subgoals (say,
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   614
those concerned with type checking).  A heuristic function might simply
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   615
count the subgoals.
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   616
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   617
\item[\ttindexbold{hyp_subst_tacs}] is a list of tactics for substitution in
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   618
the hypotheses, typically created by \ttindex{HypsubstFun} (see
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   619
Chapter~\ref{substitution}).  This list can, of course, be empty.  The
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   620
tactics are assumed to be safe!
308
f4cecad9b6a0 modifications towards final draft
lcp
parents: 286
diff changeset
   621
\end{ttdescription}
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   622
The functor is not at all sensitive to the formalization of the
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   623
object-logic.  It does not even examine the rules, but merely applies them
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   624
according to its fixed strategy.  The functor resides in {\tt
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   625
Provers/classical.ML} in the Isabelle distribution directory.
104
d8205bb279a7 Initial revision
lcp
parents:
diff changeset
   626
319
f143f7686cd6 penultimate Springer draft
lcp
parents: 308
diff changeset
   627
\index{classical reasoner|)}