| author | paulson | 
| Thu, 10 Dec 2009 17:34:09 +0000 | |
| changeset 34054 | 8e07304ecd0c | 
| parent 33057 | 764547b68538 | 
| child 42637 | 381fdcab0f36 | 
| permissions | -rw-r--r-- | 
| 10792 | 1  | 
% $Id$  | 
| 25264 | 2  | 
%!TEX root = ../tutorial.tex  | 
| 10295 | 3  | 
\chapter{The Rules of the Game}
 | 
4  | 
\label{chap:rules}
 | 
|
5  | 
||
| 11077 | 6  | 
This chapter outlines the concepts and techniques that underlie reasoning  | 
7  | 
in Isabelle. Until now, we have proved everything using only induction and  | 
|
| 13439 | 8  | 
simplification, but any serious verification project requires more elaborate  | 
| 11077 | 9  | 
forms of inference. The chapter also introduces the fundamentals of  | 
10  | 
predicate logic. The first examples in this chapter will consist of  | 
|
11  | 
detailed, low-level proof steps. Later, we shall see how to automate such  | 
|
12  | 
reasoning using the methods  | 
|
13  | 
\isa{blast},
 | 
|
14  | 
\isa{auto} and others.  Backward or goal-directed proof is our usual style,
 | 
|
15  | 
but the chapter also introduces forward reasoning, where one theorem is  | 
|
16  | 
transformed to yield another.  | 
|
| 10295 | 17  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
18  | 
\section{Natural Deduction}
 | 
| 10295 | 19  | 
|
| 11077 | 20  | 
\index{natural deduction|(}%
 | 
| 10295 | 21  | 
In Isabelle, proofs are constructed using inference rules. The  | 
| 11406 | 22  | 
most familiar inference rule is probably \emph{modus ponens}:%
 | 
23  | 
\index{modus ponens@\emph{modus ponens}} 
 | 
|
| 10295 | 24  | 
\[ \infer{Q}{P\imp Q & P} \]
 | 
| 11406 | 25  | 
This rule says that from $P\imp Q$ and $P$ we may infer~$Q$.  | 
| 10295 | 26  | 
|
| 11406 | 27  | 
\textbf{Natural deduction} is an attempt to formalize logic in a way 
 | 
| 10295 | 28  | 
that mirrors human reasoning patterns.  | 
29  | 
For each logical symbol (say, $\conj$), there  | 
|
| 11406 | 30  | 
are two kinds of rules: \textbf{introduction} and \textbf{elimination} rules. 
 | 
| 10295 | 31  | 
The introduction rules allow us to infer this symbol (say, to  | 
32  | 
infer conjunctions). The elimination rules allow us to deduce  | 
|
33  | 
consequences from this symbol. Ideally each rule should mention  | 
|
34  | 
one symbol only. For predicate logic this can be  | 
|
35  | 
done, but when users define their own concepts they typically  | 
|
| 11255 | 36  | 
have to refer to other symbols as well. It is best not to be dogmatic.  | 
| 10295 | 37  | 
|
38  | 
Natural deduction generally deserves its name. It is easy to use. Each  | 
|
39  | 
proof step consists of identifying the outermost symbol of a formula and  | 
|
40  | 
applying the corresponding rule. It creates new subgoals in  | 
|
41  | 
an obvious way from parts of the chosen formula. Expanding the  | 
|
42  | 
definitions of constants can blow up the goal enormously. Deriving natural  | 
|
43  | 
deduction rules for such constants lets us reason in terms of their key  | 
|
44  | 
properties, which might otherwise be obscured by the technicalities of its  | 
|
45  | 
definition. Natural deduction rules also lend themselves to automation.  | 
|
46  | 
Isabelle's  | 
|
| 11406 | 47  | 
\textbf{classical reasoner} accepts any suitable  collection of natural deduction
 | 
| 10295 | 48  | 
rules and uses them to search for proofs automatically. Isabelle is designed around  | 
| 11077 | 49  | 
natural deduction and many of its tools use the terminology of introduction  | 
50  | 
and elimination rules.%  | 
|
51  | 
\index{natural deduction|)}
 | 
|
| 10295 | 52  | 
|
53  | 
||
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
54  | 
\section{Introduction Rules}
 | 
| 10295 | 55  | 
|
| 11077 | 56  | 
\index{introduction rules|(}%
 | 
57  | 
An introduction rule tells us when we can infer a formula  | 
|
| 10295 | 58  | 
containing a specific logical symbol. For example, the conjunction  | 
59  | 
introduction rule says that if we have $P$ and if we have $Q$ then  | 
|
60  | 
we have $P\conj Q$. In a mathematics text, it is typically shown  | 
|
61  | 
like this:  | 
|
62  | 
\[  \infer{P\conj Q}{P & Q} \]
 | 
|
63  | 
The rule introduces the conjunction  | 
|
| 10971 | 64  | 
symbol~($\conj$) in its conclusion. In Isabelle proofs we  | 
| 10295 | 65  | 
mainly reason backwards. When we apply this rule, the subgoal already has  | 
66  | 
the form of a conjunction; the proof step makes this conjunction symbol  | 
|
67  | 
disappear.  | 
|
68  | 
||
69  | 
In Isabelle notation, the rule looks like this:  | 
|
70  | 
\begin{isabelle}
 | 
|
| 11417 | 71  | 
\isasymlbrakk?P;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P\ \isasymand\ ?Q\rulenamedx{conjI}
 | 
| 10295 | 72  | 
\end{isabelle}
 | 
73  | 
Carefully examine the syntax. The premises appear to the  | 
|
74  | 
left of the arrow and the conclusion to the right. The premises (if  | 
|
75  | 
more than one) are grouped using the fat brackets. The question marks  | 
|
| 11406 | 76  | 
indicate \textbf{schematic variables} (also called
 | 
77  | 
\textbf{unknowns}):\index{unknowns|bold} they may
 | 
|
| 10295 | 78  | 
be replaced by arbitrary formulas. If we use the rule backwards, Isabelle  | 
79  | 
tries to unify the current subgoal with the conclusion of the rule, which  | 
|
80  | 
has the form \isa{?P\ \isasymand\ ?Q}.  (Unification is discussed below,
 | 
|
| 11428 | 81  | 
{\S}\ref{sec:unification}.)  If successful,
 | 
| 10295 | 82  | 
it yields new subgoals given by the formulas assigned to  | 
83  | 
\isa{?P} and \isa{?Q}.
 | 
|
84  | 
||
| 12333 | 85  | 
The following trivial proof illustrates how rules work. It also introduces a  | 
86  | 
style of indentation. If a command adds a new subgoal, then the next  | 
|
87  | 
command's indentation is increased by one space; if it proves a subgoal, then  | 
|
88  | 
the indentation is reduced. This provides the reader with hints about the  | 
|
89  | 
subgoal structure.  | 
|
| 10295 | 90  | 
\begin{isabelle}
 | 
| 10596 | 91  | 
\isacommand{lemma}\ conj_rule:\ "\isasymlbrakk P;\
 | 
| 10295 | 92  | 
Q\isasymrbrakk\ \isasymLongrightarrow\ P\ \isasymand\  | 
| 10301 | 93  | 
(Q\ \isasymand\ P)"\isanewline  | 
| 10295 | 94  | 
\isacommand{apply}\ (rule\ conjI)\isanewline
 | 
95  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
|
96  | 
\isacommand{apply}\ (rule\ conjI)\isanewline
 | 
|
97  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
|
98  | 
\isacommand{apply}\ assumption
 | 
|
99  | 
\end{isabelle}
 | 
|
100  | 
At the start, Isabelle presents  | 
|
101  | 
us with the assumptions (\isa{P} and~\isa{Q}) and with the goal to be proved,
 | 
|
102  | 
\isa{P\ \isasymand\
 | 
|
103  | 
(Q\ \isasymand\ P)}. We are working backwards, so when we  | 
|
104  | 
apply conjunction introduction, the rule removes the outermost occurrence  | 
|
105  | 
of the \isa{\isasymand} symbol.  To apply a  rule to a subgoal, we apply
 | 
|
| 11406 | 106  | 
the proof method \isa{rule} --- here with \isa{conjI}, the  conjunction
 | 
| 10295 | 107  | 
introduction rule.  | 
108  | 
\begin{isabelle}
 | 
|
| 10596 | 109  | 
%\isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P\ \isasymand\ Q\  | 
| 10295 | 110  | 
%\isasymand\ P\isanewline  | 
| 10596 | 111  | 
\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P\isanewline  | 
112  | 
\ 2.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ Q\ \isasymand\ P  | 
|
| 10295 | 113  | 
\end{isabelle}
 | 
114  | 
Isabelle leaves two new subgoals: the two halves of the original conjunction.  | 
|
115  | 
The first is simply \isa{P}, which is trivial, since \isa{P} is among 
 | 
|
| 11406 | 116  | 
the assumptions.  We can apply the \methdx{assumption} 
 | 
| 10295 | 117  | 
method, which proves a subgoal by finding a matching assumption.  | 
118  | 
\begin{isabelle}
 | 
|
| 10596 | 119  | 
\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\  | 
| 10295 | 120  | 
Q\ \isasymand\ P  | 
121  | 
\end{isabelle}
 | 
|
122  | 
We are left with the subgoal of proving  | 
|
123  | 
\isa{Q\ \isasymand\ P} from the assumptions \isa{P} and~\isa{Q}.  We apply
 | 
|
124  | 
\isa{rule conjI} again. 
 | 
|
125  | 
\begin{isabelle}
 | 
|
| 10596 | 126  | 
\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ Q\isanewline  | 
127  | 
\ 2.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P  | 
|
| 10295 | 128  | 
\end{isabelle}
 | 
129  | 
We are left with two new subgoals, \isa{Q} and~\isa{P}, each of which can be proved
 | 
|
| 11077 | 130  | 
using the \isa{assumption} method.%
 | 
131  | 
\index{introduction rules|)}
 | 
|
| 10295 | 132  | 
|
133  | 
||
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
134  | 
\section{Elimination Rules}
 | 
| 10295 | 135  | 
|
| 11077 | 136  | 
\index{elimination rules|(}%
 | 
137  | 
Elimination rules work in the opposite direction from introduction  | 
|
| 10295 | 138  | 
rules. In the case of conjunction, there are two such rules.  | 
139  | 
From $P\conj Q$ we infer $P$. also, from $P\conj Q$  | 
|
140  | 
we infer $Q$:  | 
|
141  | 
\[ \infer{P}{P\conj Q} \qquad \infer{Q}{P\conj Q}  \]
 | 
|
142  | 
||
143  | 
Now consider disjunction. There are two introduction rules, which resemble inverted forms of the  | 
|
144  | 
conjunction elimination rules:  | 
|
145  | 
\[ \infer{P\disj Q}{P} \qquad \infer{P\disj Q}{Q}  \]
 | 
|
146  | 
||
147  | 
What is the disjunction elimination rule? The situation is rather different from  | 
|
148  | 
conjunction. From $P\disj Q$ we cannot conclude that $P$ is true and we  | 
|
149  | 
cannot conclude that $Q$ is true; there are no direct  | 
|
150  | 
elimination rules of the sort that we have seen for conjunction. Instead,  | 
|
151  | 
there is an elimination rule that works indirectly. If we are trying to prove  | 
|
152  | 
something else, say $R$, and we know that $P\disj Q$ holds, then we have to consider  | 
|
153  | 
two cases. We can assume that $P$ is true and prove $R$ and then assume that $Q$ is  | 
|
154  | 
true and prove $R$ a second time. Here we see a fundamental concept used in natural  | 
|
| 11406 | 155  | 
deduction:  that of the \textbf{assumptions}. We have to prove $R$ twice, under
 | 
| 10295 | 156  | 
different assumptions. The assumptions are local to these subproofs and are visible  | 
157  | 
nowhere else.  | 
|
158  | 
||
159  | 
In a logic text, the disjunction elimination rule might be shown  | 
|
160  | 
like this:  | 
|
161  | 
\[ \infer{R}{P\disj Q & \infer*{R}{[P]} & \infer*{R}{[Q]}} \]
 | 
|
162  | 
The assumptions $[P]$ and $[Q]$ are bracketed  | 
|
163  | 
to emphasize that they are local to their subproofs. In Isabelle  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
164  | 
notation, the already-familiar \isa{\isasymLongrightarrow} syntax serves the
 | 
| 10295 | 165  | 
same purpose:  | 
166  | 
\begin{isabelle}
 | 
|
| 11417 | 167  | 
\isasymlbrakk?P\ \isasymor\ ?Q;\ ?P\ \isasymLongrightarrow\ ?R;\ ?Q\ \isasymLongrightarrow\ ?R\isasymrbrakk\ \isasymLongrightarrow\ ?R\rulenamedx{disjE}
 | 
| 10295 | 168  | 
\end{isabelle}
 | 
169  | 
When we use this sort of elimination rule backwards, it produces  | 
|
| 10971 | 170  | 
a case split. (We have seen this before, in proofs by induction.) The following proof  | 
| 10295 | 171  | 
illustrates the use of disjunction elimination.  | 
172  | 
\begin{isabelle}
 | 
|
| 10301 | 173  | 
\isacommand{lemma}\ disj_swap:\ "P\ \isasymor\ Q\ 
 | 
| 10295 | 174  | 
\isasymLongrightarrow\ Q\ \isasymor\ P"\isanewline  | 
175  | 
\isacommand{apply}\ (erule\ disjE)\isanewline
 | 
|
176  | 
\ \isacommand{apply}\ (rule\ disjI2)\isanewline
 | 
|
177  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
|
178  | 
\isacommand{apply}\ (rule\ disjI1)\isanewline
 | 
|
179  | 
\isacommand{apply}\ assumption
 | 
|
180  | 
\end{isabelle}
 | 
|
181  | 
We assume \isa{P\ \isasymor\ Q} and
 | 
|
182  | 
must prove \isa{Q\ \isasymor\ P}\@.  Our first step uses the disjunction
 | 
|
| 11428 | 183  | 
elimination rule, \isa{disjE}\@.  We invoke it using \methdx{erule}, a
 | 
| 11406 | 184  | 
method designed to work with elimination rules. It looks for an assumption that  | 
| 11077 | 185  | 
matches the rule's first premise. It deletes the matching assumption,  | 
186  | 
regards the first premise as proved and returns subgoals corresponding to  | 
|
187  | 
the remaining premises.  When we apply \isa{erule} to \isa{disjE}, only two
 | 
|
188  | 
subgoals result.  This is better than applying it using \isa{rule}
 | 
|
189  | 
to get three subgoals, then proving the first by assumption: the other  | 
|
190  | 
subgoals would have the redundant assumption  | 
|
191  | 
\hbox{\isa{P\ \isasymor\ Q}}.
 | 
|
| 11406 | 192  | 
Most of the time, \isa{erule} is  the best way to use elimination rules, since it
 | 
193  | 
replaces an assumption by its subformulas; only rarely does the original  | 
|
194  | 
assumption remain useful.  | 
|
| 10295 | 195  | 
|
196  | 
\begin{isabelle}
 | 
|
197  | 
%P\ \isasymor\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P\isanewline  | 
|
198  | 
\ 1.\ P\ \isasymLongrightarrow\ Q\ \isasymor\ P\isanewline  | 
|
199  | 
\ 2.\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P  | 
|
200  | 
\end{isabelle}
 | 
|
| 11077 | 201  | 
These are the two subgoals returned by \isa{erule}.  The first assumes
 | 
202  | 
\isa{P} and the  second assumes \isa{Q}.  Tackling the first subgoal, we
 | 
|
203  | 
need to  show \isa{Q\ \isasymor\ P}\@.  The second introduction rule
 | 
|
204  | 
(\isa{disjI2}) can reduce this  to \isa{P}, which matches the assumption.
 | 
|
205  | 
So, we apply the  | 
|
| 10596 | 206  | 
\isa{rule}  method with \isa{disjI2} \ldots
 | 
| 10295 | 207  | 
\begin{isabelle}
 | 
208  | 
\ 1.\ P\ \isasymLongrightarrow\ P\isanewline  | 
|
209  | 
\ 2.\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P  | 
|
210  | 
\end{isabelle}
 | 
|
| 10596 | 211  | 
\ldots and finish off with the \isa{assumption} 
 | 
| 10295 | 212  | 
method. We are left with the other subgoal, which  | 
213  | 
assumes \isa{Q}.  
 | 
|
214  | 
\begin{isabelle}
 | 
|
215  | 
\ 1.\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P  | 
|
216  | 
\end{isabelle}
 | 
|
217  | 
Its proof is similar, using the introduction  | 
|
218  | 
rule \isa{disjI1}. 
 | 
|
219  | 
||
220  | 
The result of this proof is a new inference rule \isa{disj_swap}, which is neither 
 | 
|
221  | 
an introduction nor an elimination rule, but which might  | 
|
222  | 
be useful. We can use it to replace any goal of the form $Q\disj P$  | 
|
| 27167 | 223  | 
by one of the form $P\disj Q$.%  | 
| 11077 | 224  | 
\index{elimination rules|)}
 | 
| 10295 | 225  | 
|
226  | 
||
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
227  | 
\section{Destruction Rules: Some Examples}
 | 
| 10295 | 228  | 
|
| 11077 | 229  | 
\index{destruction rules|(}%
 | 
| 10295 | 230  | 
Now let us examine the analogous proof for conjunction.  | 
231  | 
\begin{isabelle}
 | 
|
| 10301 | 232  | 
\isacommand{lemma}\ conj_swap:\ "P\ \isasymand\ Q\ \isasymLongrightarrow\ Q\ \isasymand\ P"\isanewline
 | 
| 10295 | 233  | 
\isacommand{apply}\ (rule\ conjI)\isanewline
 | 
234  | 
\ \isacommand{apply}\ (drule\ conjunct2)\isanewline
 | 
|
235  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
|
236  | 
\isacommand{apply}\ (drule\ conjunct1)\isanewline
 | 
|
237  | 
\isacommand{apply}\ assumption
 | 
|
238  | 
\end{isabelle}
 | 
|
239  | 
Recall that the conjunction elimination rules --- whose Isabelle names are  | 
|
240  | 
\isa{conjunct1} and \isa{conjunct2} --- simply return the first or second half
 | 
|
241  | 
of a conjunction. Rules of this sort (where the conclusion is a subformula of a  | 
|
| 11406 | 242  | 
premise) are called \textbf{destruction} rules because they take apart and destroy
 | 
| 10978 | 243  | 
a premise.%  | 
| 10295 | 244  | 
\footnote{This Isabelle terminology has no counterpart in standard logic texts, 
 | 
245  | 
although the distinction between the two forms of elimination rule is well known.  | 
|
| 11406 | 246  | 
Girard \cite[page 74]{girard89},\index{Girard, Jean-Yves|fnote}
 | 
247  | 
for example, writes ``The elimination rules  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
248  | 
[for $\disj$ and $\exists$] are very  | 
| 10295 | 249  | 
bad. What is catastrophic about them is the parasitic presence of a formula [$R$]  | 
250  | 
which has no structural link with the formula which is eliminated.''}  | 
|
251  | 
||
252  | 
The first proof step applies conjunction introduction, leaving  | 
|
253  | 
two subgoals:  | 
|
254  | 
\begin{isabelle}
 | 
|
255  | 
%P\ \isasymand\ Q\ \isasymLongrightarrow\ Q\ \isasymand\ P\isanewline  | 
|
256  | 
\ 1.\ P\ \isasymand\ Q\ \isasymLongrightarrow\ Q\isanewline  | 
|
257  | 
\ 2.\ P\ \isasymand\ Q\ \isasymLongrightarrow\ P  | 
|
258  | 
\end{isabelle}
 | 
|
259  | 
||
260  | 
To invoke the elimination rule, we apply a new method, \isa{drule}. 
 | 
|
| 11406 | 261  | 
Think of the \isa{d} as standing for \textbf{destruction} (or \textbf{direct}, if
 | 
| 10295 | 262  | 
you prefer). Applying the  | 
263  | 
second conjunction rule using \isa{drule} replaces the assumption 
 | 
|
264  | 
\isa{P\ \isasymand\ Q} by \isa{Q}. 
 | 
|
265  | 
\begin{isabelle}
 | 
|
266  | 
\ 1.\ Q\ \isasymLongrightarrow\ Q\isanewline  | 
|
267  | 
\ 2.\ P\ \isasymand\ Q\ \isasymLongrightarrow\ P  | 
|
268  | 
\end{isabelle}
 | 
|
269  | 
The resulting subgoal can be proved by applying \isa{assumption}.
 | 
|
270  | 
The other subgoal is similarly proved, using the \isa{conjunct1} rule and the 
 | 
|
271  | 
\isa{assumption} method.
 | 
|
272  | 
||
273  | 
Choosing among the methods \isa{rule}, \isa{erule} and \isa{drule} is up to 
 | 
|
274  | 
you. Isabelle does not attempt to work out whether a rule  | 
|
275  | 
is an introduction rule or an elimination rule. The  | 
|
276  | 
method determines how the rule will be interpreted. Many rules  | 
|
277  | 
can be used in more than one way.  For example, \isa{disj_swap} can 
 | 
|
278  | 
be applied to assumptions as well as to goals; it replaces any  | 
|
279  | 
assumption of the form  | 
|
280  | 
$P\disj Q$ by a one of the form $Q\disj P$.  | 
|
281  | 
||
282  | 
Destruction rules are simpler in form than indirect rules such as \isa{disjE},
 | 
|
283  | 
but they can be inconvenient. Each of the conjunction rules discards half  | 
|
284  | 
of the formula, when usually we want to take both parts of the conjunction as new  | 
|
285  | 
assumptions. The easiest way to do so is by using an  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
286  | 
alternative conjunction elimination rule that resembles \isa{disjE}\@.  It is
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
287  | 
seldom, if ever, seen in logic books. In Isabelle syntax it looks like this:  | 
| 10295 | 288  | 
\begin{isabelle}
 | 
| 11417 | 289  | 
\isasymlbrakk?P\ \isasymand\ ?Q;\ \isasymlbrakk?P;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?R\isasymrbrakk\ \isasymLongrightarrow\ ?R\rulenamedx{conjE}
 | 
| 10295 | 290  | 
\end{isabelle}
 | 
| 11077 | 291  | 
\index{destruction rules|)}
 | 
| 10295 | 292  | 
|
293  | 
\begin{exercise}
 | 
|
| 11077 | 294  | 
Use the rule \isa{conjE} to shorten the proof above. 
 | 
| 10295 | 295  | 
\end{exercise}
 | 
296  | 
||
297  | 
||
298  | 
\section{Implication}
 | 
|
299  | 
||
| 11077 | 300  | 
\index{implication|(}%
 | 
| 11406 | 301  | 
At the start of this chapter, we saw the rule \emph{modus ponens}.  It is, in fact,
 | 
| 10295 | 302  | 
a destruction rule. The matching introduction rule looks like this  | 
303  | 
in Isabelle:  | 
|
304  | 
\begin{isabelle}
 | 
|
305  | 
(?P\ \isasymLongrightarrow\ ?Q)\ \isasymLongrightarrow\ ?P\  | 
|
| 11417 | 306  | 
\isasymlongrightarrow\ ?Q\rulenamedx{impI}
 | 
| 10295 | 307  | 
\end{isabelle}
 | 
| 12535 | 308  | 
And this is \emph{modus ponens}\index{modus ponens@\emph{modus ponens}}:
 | 
| 10295 | 309  | 
\begin{isabelle}
 | 
310  | 
\isasymlbrakk?P\ \isasymlongrightarrow\ ?Q;\ ?P\isasymrbrakk\  | 
|
311  | 
\isasymLongrightarrow\ ?Q  | 
|
| 11417 | 312  | 
\rulenamedx{mp}
 | 
| 10295 | 313  | 
\end{isabelle}
 | 
314  | 
||
| 11077 | 315  | 
Here is a proof using the implication rules. This  | 
| 10295 | 316  | 
lemma performs a sort of uncurrying, replacing the two antecedents  | 
| 11077 | 317  | 
of a nested implication by a conjunction. The proof illustrates  | 
318  | 
how assumptions work. At each proof step, the subgoals inherit the previous  | 
|
319  | 
assumptions, perhaps with additions or deletions. Rules such as  | 
|
320  | 
\isa{impI} and \isa{disjE} add assumptions, while applying \isa{erule} or
 | 
|
321  | 
\isa{drule} deletes the matching assumption.
 | 
|
| 10295 | 322  | 
\begin{isabelle}
 | 
323  | 
\isacommand{lemma}\ imp_uncurry:\
 | 
|
| 10301 | 324  | 
"P\ \isasymlongrightarrow\ (Q\  | 
| 10295 | 325  | 
\isasymlongrightarrow\ R)\ \isasymLongrightarrow\ P\  | 
326  | 
\isasymand\ Q\ \isasymlongrightarrow\  | 
|
327  | 
R"\isanewline  | 
|
328  | 
\isacommand{apply}\ (rule\ impI)\isanewline
 | 
|
329  | 
\isacommand{apply}\ (erule\ conjE)\isanewline
 | 
|
330  | 
\isacommand{apply}\ (drule\ mp)\isanewline
 | 
|
331  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
|
332  | 
\isacommand{apply}\ (drule\ mp)\isanewline
 | 
|
333  | 
\ \ \isacommand{apply}\ assumption\isanewline
 | 
|
334  | 
\ \isacommand{apply}\ assumption
 | 
|
335  | 
\end{isabelle}
 | 
|
336  | 
First, we state the lemma and apply implication introduction (\isa{rule impI}), 
 | 
|
337  | 
which moves the conjunction to the assumptions.  | 
|
338  | 
\begin{isabelle}
 | 
|
339  | 
%P\ \isasymlongrightarrow\ Q\ \isasymlongrightarrow\ R\ \isasymLongrightarrow\ P\  | 
|
340  | 
%\isasymand\ Q\ \isasymlongrightarrow\ R\isanewline  | 
|
| 10596 | 341  | 
\ 1.\ \isasymlbrakk P\ \isasymlongrightarrow\ Q\ \isasymlongrightarrow\ R;\ P\ \isasymand\ Q\isasymrbrakk\ \isasymLongrightarrow\ R  | 
| 10295 | 342  | 
\end{isabelle}
 | 
343  | 
Next, we apply conjunction elimination (\isa{erule conjE}), which splits this
 | 
|
344  | 
conjunction into two parts.  | 
|
345  | 
\begin{isabelle}
 | 
|
| 10596 | 346  | 
\ 1.\ \isasymlbrakk P\ \isasymlongrightarrow\ Q\ \isasymlongrightarrow\ R;\ P;\  | 
| 10295 | 347  | 
Q\isasymrbrakk\ \isasymLongrightarrow\ R  | 
348  | 
\end{isabelle}
 | 
|
349  | 
Now, we work on the assumption \isa{P\ \isasymlongrightarrow\ (Q\
 | 
|
350  | 
\isasymlongrightarrow\ R)}, where the parentheses have been inserted for  | 
|
351  | 
clarity. The nested implication requires two applications of  | 
|
352  | 
\textit{modus ponens}: \isa{drule mp}.  The first use  yields the
 | 
|
353  | 
implication \isa{Q\
 | 
|
354  | 
\isasymlongrightarrow\ R}, but first we must prove the extra subgoal  | 
|
355  | 
\isa{P}, which we do by assumption. 
 | 
|
356  | 
\begin{isabelle}
 | 
|
| 10596 | 357  | 
\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P\isanewline  | 
358  | 
\ 2.\ \isasymlbrakk P;\ Q;\ Q\ \isasymlongrightarrow\ R\isasymrbrakk\ \isasymLongrightarrow\ R  | 
|
| 10295 | 359  | 
\end{isabelle}
 | 
360  | 
Repeating these steps for \isa{Q\
 | 
|
361  | 
\isasymlongrightarrow\ R} yields the conclusion we seek, namely~\isa{R}.
 | 
|
362  | 
\begin{isabelle}
 | 
|
| 10596 | 363  | 
\ 1.\ \isasymlbrakk P;\ Q;\ Q\ \isasymlongrightarrow\ R\isasymrbrakk\  | 
| 10295 | 364  | 
\isasymLongrightarrow\ R  | 
365  | 
\end{isabelle}
 | 
|
366  | 
||
367  | 
The symbols \isa{\isasymLongrightarrow} and \isa{\isasymlongrightarrow}
 | 
|
368  | 
both stand for implication, but they differ in many respects. Isabelle  | 
|
369  | 
uses \isa{\isasymLongrightarrow} to express inference rules; the symbol is
 | 
|
370  | 
built-in and Isabelle's inference mechanisms treat it specially. On the  | 
|
371  | 
other hand, \isa{\isasymlongrightarrow} is just one of the many connectives
 | 
|
372  | 
available in higher-order logic. We reason about it using inference rules  | 
|
373  | 
such as \isa{impI} and \isa{mp}, just as we reason about the other
 | 
|
374  | 
connectives.  You will have to use \isa{\isasymlongrightarrow} in any
 | 
|
375  | 
context that requires a formula of higher-order logic. Use  | 
|
376  | 
\isa{\isasymLongrightarrow} to separate a theorem's preconditions from its
 | 
|
| 11077 | 377  | 
conclusion.%  | 
378  | 
\index{implication|)}
 | 
|
| 10295 | 379  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
380  | 
\medskip  | 
| 11406 | 381  | 
\index{by@\isacommand{by} (command)|(}%
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
382  | 
The \isacommand{by} command is useful for proofs like these that use
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
383  | 
\isa{assumption} heavily.  It executes an
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
384  | 
\isacommand{apply} command, then tries to prove all remaining subgoals using
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
385  | 
\isa{assumption}.  Since (if successful) it ends the proof, it also replaces the 
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
386  | 
\isacommand{done} symbol.  For example, the proof above can be shortened:
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
387  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
388  | 
\isacommand{lemma}\ imp_uncurry:\
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
389  | 
"P\ \isasymlongrightarrow\ (Q\  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
390  | 
\isasymlongrightarrow\ R)\ \isasymLongrightarrow\ P\  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
391  | 
\isasymand\ Q\ \isasymlongrightarrow\  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
392  | 
R"\isanewline  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
393  | 
\isacommand{apply}\ (rule\ impI)\isanewline
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
394  | 
\isacommand{apply}\ (erule\ conjE)\isanewline
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
395  | 
\isacommand{apply}\ (drule\ mp)\isanewline
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
396  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
397  | 
\isacommand{by}\ (drule\ mp)
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
398  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
399  | 
We could use \isacommand{by} to replace the final \isacommand{apply} and
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
400  | 
\isacommand{done} in any proof, but typically we use it
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
401  | 
to eliminate calls to \isa{assumption}.  It is also a nice way of expressing a
 | 
| 11406 | 402  | 
one-line proof.%  | 
403  | 
\index{by@\isacommand{by} (command)|)}
 | 
|
404  | 
||
| 10295 | 405  | 
|
406  | 
||
407  | 
\section{Negation}
 | 
|
408  | 
||
| 11077 | 409  | 
\index{negation|(}%
 | 
| 10295 | 410  | 
Negation causes surprising complexity in proofs. Its natural  | 
411  | 
deduction rules are straightforward, but additional rules seem  | 
|
| 11077 | 412  | 
necessary in order to handle negated assumptions gracefully. This section  | 
413  | 
also illustrates the \isa{intro} method: a convenient way of
 | 
|
414  | 
applying introduction rules.  | 
|
| 10295 | 415  | 
|
| 11428 | 416  | 
Negation introduction deduces $\lnot P$ if assuming $P$ leads to a  | 
| 10295 | 417  | 
contradiction. Negation elimination deduces any formula in the  | 
| 11428 | 418  | 
presence of $\lnot P$ together with~$P$:  | 
| 10295 | 419  | 
\begin{isabelle}
 | 
420  | 
(?P\ \isasymLongrightarrow\ False)\ \isasymLongrightarrow\ \isasymnot\ ?P%  | 
|
| 11417 | 421  | 
\rulenamedx{notI}\isanewline
 | 
| 10295 | 422  | 
\isasymlbrakk{\isasymnot}\ ?P;\ ?P\isasymrbrakk\ \isasymLongrightarrow\ ?R%
 | 
| 11417 | 423  | 
\rulenamedx{notE}
 | 
| 10295 | 424  | 
\end{isabelle}
 | 
425  | 
%  | 
|
| 11428 | 426  | 
Classical logic allows us to assume $\lnot P$  | 
| 10295 | 427  | 
when attempting to prove~$P$:  | 
428  | 
\begin{isabelle}
 | 
|
429  | 
(\isasymnot\ ?P\ \isasymLongrightarrow\ ?P)\ \isasymLongrightarrow\ ?P%  | 
|
| 11417 | 430  | 
\rulenamedx{classical}
 | 
| 10295 | 431  | 
\end{isabelle}
 | 
| 11077 | 432  | 
|
| 11406 | 433  | 
\index{contrapositives|(}%
 | 
| 11428 | 434  | 
The implications $P\imp Q$ and $\lnot Q\imp\lnot P$ are logically  | 
| 11077 | 435  | 
equivalent, and each is called the  | 
| 11406 | 436  | 
\textbf{contrapositive} of the other.  Four further rules support
 | 
| 11077 | 437  | 
reasoning about contrapositives. They differ in the placement of the  | 
438  | 
negation symbols:  | 
|
| 10295 | 439  | 
\begin{isabelle}
 | 
440  | 
\isasymlbrakk?Q;\ \isasymnot\ ?P\ \isasymLongrightarrow\ \isasymnot\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%  | 
|
441  | 
\rulename{contrapos_pp}\isanewline
 | 
|
| 11406 | 442  | 
\isasymlbrakk?Q;\ ?P\ \isasymLongrightarrow\ \isasymnot\ ?Q\isasymrbrakk\ \isasymLongrightarrow\  | 
443  | 
\isasymnot\ ?P%  | 
|
444  | 
\rulename{contrapos_pn}\isanewline
 | 
|
| 10295 | 445  | 
\isasymlbrakk{\isasymnot}\ ?Q;\ \isasymnot\ ?P\ \isasymLongrightarrow\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%
 | 
446  | 
\rulename{contrapos_np}\isanewline
 | 
|
447  | 
\isasymlbrakk{\isasymnot}\ ?Q;\ ?P\ \isasymLongrightarrow\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ \isasymnot\ ?P%
 | 
|
448  | 
\rulename{contrapos_nn}
 | 
|
449  | 
\end{isabelle}
 | 
|
450  | 
%  | 
|
| 11077 | 451  | 
These rules are typically applied using the \isa{erule} method, where 
 | 
| 10295 | 452  | 
their effect is to form a contrapositive from an  | 
| 11406 | 453  | 
assumption and the goal's conclusion.%  | 
454  | 
\index{contrapositives|)}
 | 
|
| 10295 | 455  | 
|
456  | 
The most important of these is \isa{contrapos_np}.  It is useful
 | 
|
457  | 
for applying introduction rules to negated assumptions. For instance,  | 
|
| 11428 | 458  | 
the assumption $\lnot(P\imp Q)$ is equivalent to the conclusion $P\imp Q$ and we  | 
| 10295 | 459  | 
might want to use conjunction introduction on it.  | 
460  | 
Before we can do so, we must move that assumption so that it  | 
|
461  | 
becomes the conclusion. The following proof demonstrates this  | 
|
462  | 
technique:  | 
|
463  | 
\begin{isabelle}
 | 
|
464  | 
\isacommand{lemma}\ "\isasymlbrakk{\isasymnot}(P{\isasymlongrightarrow}Q);\
 | 
|
465  | 
\isasymnot(R{\isasymlongrightarrow}Q)\isasymrbrakk\ \isasymLongrightarrow\
 | 
|
466  | 
R"\isanewline  | 
|
| 10971 | 467  | 
\isacommand{apply}\ (erule_tac\ Q = "R{\isasymlongrightarrow}Q"\ \isakeyword{in}\
 | 
| 10295 | 468  | 
contrapos_np)\isanewline  | 
| 12408 | 469  | 
\isacommand{apply}\ (intro\ impI)\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
470  | 
\isacommand{by}\ (erule\ notE)
 | 
| 10295 | 471  | 
\end{isabelle}
 | 
472  | 
%  | 
|
473  | 
There are two negated assumptions and we need to exchange the conclusion with the  | 
|
474  | 
second one.  The method \isa{erule contrapos_np} would select the first assumption,
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
475  | 
which we do not want. So we specify the desired assumption explicitly  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
476  | 
using a new method, \isa{erule_tac}.  This is the resulting subgoal: 
 | 
| 10295 | 477  | 
\begin{isabelle}
 | 
478  | 
\ 1.\ \isasymlbrakk{\isasymnot}\ (P\ \isasymlongrightarrow\ Q);\ \isasymnot\
 | 
|
479  | 
R\isasymrbrakk\ \isasymLongrightarrow\ R\ \isasymlongrightarrow\ Q%  | 
|
480  | 
\end{isabelle}
 | 
|
481  | 
The former conclusion, namely \isa{R}, now appears negated among the assumptions,
 | 
|
482  | 
while the negated formula \isa{R\ \isasymlongrightarrow\ Q} becomes the new
 | 
|
483  | 
conclusion.  | 
|
484  | 
||
| 11406 | 485  | 
We can now apply introduction rules.  We use the \methdx{intro} method, which
 | 
| 12408 | 486  | 
repeatedly applies the given introduction rules. Here its effect is equivalent  | 
| 10596 | 487  | 
to \isa{rule impI}.
 | 
488  | 
\begin{isabelle}
 | 
|
| 10295 | 489  | 
\ 1.\ \isasymlbrakk{\isasymnot}\ (P\ \isasymlongrightarrow\ Q);\ \isasymnot\ R;\
 | 
490  | 
R\isasymrbrakk\ \isasymLongrightarrow\ Q%  | 
|
491  | 
\end{isabelle}
 | 
|
492  | 
We can see a contradiction in the form of assumptions \isa{\isasymnot\ R}
 | 
|
493  | 
and~\isa{R}, which suggests using negation elimination.  If applied on its own,
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
494  | 
\isa{notE} will select the first negated assumption, which is useless.  
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
495  | 
Instead, we invoke the rule using the  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
496  | 
\isa{by} command.
 | 
| 10295 | 497  | 
Now when Isabelle selects the first assumption, it tries to prove \isa{P\
 | 
498  | 
\isasymlongrightarrow\ Q} and fails; it then backtracks, finds the  | 
|
| 10971 | 499  | 
assumption \isa{\isasymnot~R} and finally proves \isa{R} by assumption.  That
 | 
| 10295 | 500  | 
concludes the proof.  | 
501  | 
||
502  | 
\medskip  | 
|
503  | 
||
| 11077 | 504  | 
The following example may be skipped on a first reading. It involves a  | 
505  | 
peculiar but important rule, a form of disjunction introduction:  | 
|
506  | 
\begin{isabelle}
 | 
|
507  | 
(\isasymnot \ ?Q\ \isasymLongrightarrow \ ?P)\ \isasymLongrightarrow \ ?P\ \isasymor \ ?Q%  | 
|
| 11417 | 508  | 
\rulenamedx{disjCI}
 | 
| 11077 | 509  | 
\end{isabelle}
 | 
510  | 
This rule combines the effects of \isa{disjI1} and \isa{disjI2}.  Its great
 | 
|
511  | 
advantage is that we can remove the disjunction symbol without deciding  | 
|
| 11406 | 512  | 
which disjunction to prove. This treatment of disjunction is standard in sequent  | 
513  | 
and tableau calculi.  | 
|
| 11077 | 514  | 
|
| 10295 | 515  | 
\begin{isabelle}
 | 
516  | 
\isacommand{lemma}\ "(P\ \isasymor\ Q)\ \isasymand\ R\
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
517  | 
\isasymLongrightarrow\ P\ \isasymor\ (Q\ \isasymand\ R)"\isanewline  | 
| 27167 | 518  | 
\isacommand{apply}\ (rule\ disjCI)\isanewline
 | 
| 10295 | 519  | 
\isacommand{apply}\ (elim\ conjE\ disjE)\isanewline
 | 
520  | 
\ \isacommand{apply}\ assumption
 | 
|
521  | 
\isanewline  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
522  | 
\isacommand{by}\ (erule\ contrapos_np,\ rule\ conjI)
 | 
| 10295 | 523  | 
\end{isabelle}
 | 
524  | 
%  | 
|
| 27167 | 525  | 
The first proof step to applies the introduction rules \isa{disjCI}.
 | 
526  | 
The resulting subgoal has the negative assumption  | 
|
| 11077 | 527  | 
\hbox{\isa{\isasymnot(Q\ \isasymand\ R)}}.  
 | 
528  | 
||
| 10295 | 529  | 
\begin{isabelle}
 | 
530  | 
\ 1.\ \isasymlbrakk(P\ \isasymor\ Q)\ \isasymand\ R;\ \isasymnot\ (Q\ \isasymand\  | 
|
531  | 
R)\isasymrbrakk\ \isasymLongrightarrow\ P%  | 
|
532  | 
\end{isabelle}
 | 
|
| 11077 | 533  | 
Next we apply the \isa{elim} method, which repeatedly applies 
 | 
| 10295 | 534  | 
elimination rules; here, the elimination rules given  | 
| 10971 | 535  | 
in the command.  One of the subgoals is trivial (\isa{\isacommand{apply} assumption}),
 | 
536  | 
leaving us with one other:  | 
|
| 10295 | 537  | 
\begin{isabelle}
 | 
538  | 
\ 1.\ \isasymlbrakk{\isasymnot}\ (Q\ \isasymand\ R);\ R;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P%
 | 
|
539  | 
\end{isabelle}
 | 
|
540  | 
%  | 
|
541  | 
Now we must move the formula \isa{Q\ \isasymand\ R} to be the conclusion.  The
 | 
|
542  | 
combination  | 
|
543  | 
\begin{isabelle}
 | 
|
544  | 
\ \ \ \ \ (erule\ contrapos_np,\ rule\ conjI)  | 
|
545  | 
\end{isabelle}
 | 
|
546  | 
is robust: the \isa{conjI} forces the \isa{erule} to select a
 | 
|
| 10301 | 547  | 
conjunction. The two subgoals are the ones we would expect from applying  | 
| 10295 | 548  | 
conjunction introduction to  | 
| 10971 | 549  | 
\isa{Q~\isasymand~R}:  
 | 
| 10295 | 550  | 
\begin{isabelle}
 | 
| 10596 | 551  | 
\ 1.\ \isasymlbrakk R;\ Q;\ \isasymnot\ P\isasymrbrakk\ \isasymLongrightarrow\  | 
| 10295 | 552  | 
Q\isanewline  | 
| 10596 | 553  | 
\ 2.\ \isasymlbrakk R;\ Q;\ \isasymnot\ P\isasymrbrakk\ \isasymLongrightarrow\ R%  | 
| 10295 | 554  | 
\end{isabelle}
 | 
| 11077 | 555  | 
They are proved by assumption, which is implicit in the \isacommand{by}
 | 
556  | 
command.%  | 
|
557  | 
\index{negation|)}
 | 
|
558  | 
||
559  | 
||
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
560  | 
\section{Interlude: the Basic Methods for Rules}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
561  | 
|
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
562  | 
We have seen examples of many tactics that operate on individual rules. It  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
563  | 
may be helpful to review how they work given an arbitrary rule such as this:  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
564  | 
\[ \infer{Q}{P@1 & \ldots & P@n} \]
 | 
| 11406 | 565  | 
Below, we refer to $P@1$ as the \bfindex{major premise}.  This concept
 | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
566  | 
applies only to elimination and destruction rules. These rules act upon an  | 
| 11406 | 567  | 
instance of their major premise, typically to replace it by subformulas of itself.  | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
568  | 
|
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
569  | 
Suppose that the rule above is called~\isa{R}\@.  Here are the basic rule
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
570  | 
methods, most of which we have already seen:  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
571  | 
\begin{itemize}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
572  | 
\item  | 
| 11406 | 573  | 
Method \isa{rule\ R} unifies~$Q$ with the current subgoal, replacing it
 | 
574  | 
by $n$ new subgoals: instances of $P@1$, \ldots,~$P@n$.  | 
|
575  | 
This is backward reasoning and is appropriate for introduction rules.  | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
576  | 
\item  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
577  | 
Method \isa{erule\ R} unifies~$Q$ with the current subgoal and
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
578  | 
simultaneously unifies $P@1$ with some assumption. The subgoal is  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
579  | 
replaced by the $n-1$ new subgoals of proving  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
580  | 
instances of $P@2$,  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
581  | 
\ldots,~$P@n$, with the matching assumption deleted. It is appropriate for  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
582  | 
elimination rules. The method  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
583  | 
\isa{(rule\ R,\ assumption)} is similar, but it does not delete an
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
584  | 
assumption.  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
585  | 
\item  | 
| 11406 | 586  | 
Method \isa{drule\ R} unifies $P@1$ with some assumption, which it
 | 
587  | 
then deletes. The subgoal is  | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
588  | 
replaced by the $n-1$ new subgoals of proving $P@2$, \ldots,~$P@n$; an  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
589  | 
$n$th subgoal is like the original one but has an additional assumption: an  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
590  | 
instance of~$Q$. It is appropriate for destruction rules.  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
591  | 
\item  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
592  | 
Method \isa{frule\ R} is like \isa{drule\ R} except that the matching
 | 
| 11428 | 593  | 
assumption is not deleted.  (See {\S}\ref{sec:frule} below.)
 | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
594  | 
\end{itemize}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
595  | 
|
| 11406 | 596  | 
Other methods apply a rule while constraining some of its  | 
597  | 
variables. The typical form is  | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
598  | 
\begin{isabelle}
 | 
| 11406 | 599  | 
\ \ \ \ \ \methdx{rule_tac}\ $v@1$ = $t@1$ \isakeyword{and} \ldots \isakeyword{and}
 | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
600  | 
$v@k$ =  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
601  | 
$t@k$ \isakeyword{in} R
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
602  | 
\end{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
603  | 
This method behaves like \isa{rule R}, while instantiating the variables
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
604  | 
$v@1$, \ldots,  | 
| 11406 | 605  | 
$v@k$ as specified.  We similarly have \methdx{erule_tac}, \methdx{drule_tac} and
 | 
606  | 
\methdx{frule_tac}.  These methods also let us specify which subgoal to
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
607  | 
operate on. By default it is the first subgoal, as with nearly all  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
608  | 
methods, but we can specify that rule \isa{R} should be applied to subgoal
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
609  | 
number~$i$:  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
610  | 
\begin{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
611  | 
\ \ \ \ \ rule_tac\ [$i$] R  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
612  | 
\end{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
613  | 
|
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
614  | 
|
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
615  | 
|
| 11077 | 616  | 
\section{Unification and Substitution}\label{sec:unification}
 | 
617  | 
||
618  | 
\index{unification|(}%
 | 
|
| 11406 | 619  | 
As we have seen, Isabelle rules involve schematic variables, which begin with  | 
| 11077 | 620  | 
a question mark and act as  | 
| 13751 | 621  | 
placeholders for terms.  \textbf{Unification} --- well known to Prolog programmers --- is the act of
 | 
622  | 
making two terms identical, possibly replacing their schematic variables by  | 
|
| 11406 | 623  | 
terms. The simplest case is when the two terms are already the same. Next  | 
624  | 
simplest is \textbf{pattern-matching}, which replaces variables in only one of the
 | 
|
625  | 
terms. The  | 
|
| 11077 | 626  | 
\isa{rule} method typically  matches the rule's conclusion
 | 
| 13751 | 627  | 
against the current subgoal. The  | 
628  | 
\isa{assumption} method matches the current subgoal's conclusion
 | 
|
629  | 
against each of its assumptions.   Unification can instantiate variables in both terms; the \isa{rule} method can do this if the goal
 | 
|
| 11077 | 630  | 
itself contains schematic variables. Other occurrences of the variables in  | 
631  | 
the rule or proof state are updated at the same time.  | 
|
632  | 
||
633  | 
Schematic variables in goals represent unknown terms. Given a goal such  | 
|
634  | 
as $\exists x.\,P$, they let us proceed with a proof. They can be  | 
|
635  | 
filled in later, sometimes in stages and often automatically.  | 
|
636  | 
||
| 16359 | 637  | 
\begin{pgnote}
 | 
| 16523 | 638  | 
If unification fails when you think it should succeed, try setting the Proof General flag \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$
 | 
639  | 
\pgmenu{Trace Unification},
 | 
|
| 16359 | 640  | 
which makes Isabelle show the cause of unification failures (in Proof  | 
| 16523 | 641  | 
General's \pgmenu{Trace} buffer).
 | 
| 16359 | 642  | 
\end{pgnote}
 | 
| 16412 | 643  | 
\noindent  | 
| 16359 | 644  | 
For example, suppose we are trying to prove this subgoal by assumption:  | 
| 13751 | 645  | 
\begin{isabelle}
 | 
646  | 
\ 1.\ P\ (a,\ f\ (b,\ g\ (e,\ a),\ b),\ a)\ \isasymLongrightarrow \ P\ (a,\ f\ (b,\ g\ (c,\ a),\ b),\ a)  | 
|
647  | 
\end{isabelle}
 | 
|
648  | 
The \isa{assumption} method having failed, we try again with the flag set:
 | 
|
649  | 
\begin{isabelle}
 | 
|
650  | 
\isacommand{apply} assumption
 | 
|
651  | 
\end{isabelle}
 | 
|
| 16412 | 652  | 
In this trivial case, the output clearly shows that \isa{e} clashes with \isa{c}:
 | 
| 13751 | 653  | 
\begin{isabelle}
 | 
| 16412 | 654  | 
Clash: e =/= c  | 
| 13751 | 655  | 
\end{isabelle}
 | 
656  | 
||
657  | 
Isabelle uses  | 
|
| 11406 | 658  | 
\textbf{higher-order} unification, which works in the
 | 
| 13751 | 659  | 
typed $\lambda$-calculus. The procedure requires search and is potentially  | 
660  | 
undecidable. For our purposes, however, the differences from ordinary  | 
|
661  | 
unification are straightforward. It handles bound variables  | 
|
662  | 
correctly, avoiding capture. The two terms  | 
|
663  | 
\isa{{\isasymlambda}x.\ f(x,z)} and \isa{{\isasymlambda}y.\ f(y,z)} are
 | 
|
664  | 
trivially unifiable because they differ only by a bound variable renaming.  The two terms \isa{{\isasymlambda}x.\ ?P} and
 | 
|
| 11077 | 665  | 
\isa{{\isasymlambda}x.\ t x}  are not unifiable; replacing \isa{?P} by
 | 
666  | 
\isa{t x} is forbidden because the free occurrence of~\isa{x} would become
 | 
|
| 13751 | 667  | 
bound.  Unfortunately, even if \isa{trace_unify_fail} is set, Isabelle displays no information about this type of failure.
 | 
| 11406 | 668  | 
|
| 11077 | 669  | 
\begin{warn}
 | 
670  | 
Higher-order unification sometimes must invent  | 
|
671  | 
$\lambda$-terms to replace function variables,  | 
|
672  | 
which can lead to a combinatorial explosion. However, Isabelle proofs tend  | 
|
673  | 
to involve easy cases where there are few possibilities for the  | 
|
674  | 
$\lambda$-term being constructed. In the easiest case, the  | 
|
675  | 
function variable is applied only to bound variables,  | 
|
676  | 
as when we try to unify \isa{{\isasymlambda}x\ y.\ f(?h x y)} and
 | 
|
677  | 
\isa{{\isasymlambda}x\ y.\ f(x+y+a)}.  The only solution is to replace
 | 
|
678  | 
\isa{?h} by \isa{{\isasymlambda}x\ y.\ x+y+a}.  Such cases admit at most
 | 
|
679  | 
one unifier, like ordinary unification. A harder case is  | 
|
680  | 
unifying \isa{?h a} with~\isa{a+b}; it admits two solutions for \isa{?h},
 | 
|
681  | 
namely \isa{{\isasymlambda}x.~a+b} and \isa{{\isasymlambda}x.~x+b}. 
 | 
|
682  | 
Unifying \isa{?h a} with~\isa{a+a+b} admits four solutions; their number is
 | 
|
683  | 
exponential in the number of occurrences of~\isa{a} in the second term.
 | 
|
684  | 
\end{warn}
 | 
|
685  | 
||
686  | 
||
| 11406 | 687  | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
688  | 
\subsection{Substitution and the {\tt\slshape subst} Method}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
689  | 
\label{sec:subst}
 | 
| 11077 | 690  | 
|
691  | 
\index{substitution|(}%
 | 
|
| 11406 | 692  | 
Isabelle also uses function variables to express \textbf{substitution}. 
 | 
| 11077 | 693  | 
A typical substitution rule allows us to replace one term by  | 
694  | 
another if we know that two terms are equal.  | 
|
695  | 
\[ \infer{P[t/x]}{s=t & P[s/x]} \]
 | 
|
696  | 
The rule uses a notation for substitution: $P[t/x]$ is the result of  | 
|
697  | 
replacing $x$ by~$t$ in~$P$. The rule only substitutes in the positions  | 
|
698  | 
designated by~$x$. For example, it can  | 
|
699  | 
derive symmetry of equality from reflexivity. Using $x=s$ for~$P$  | 
|
700  | 
replaces just the first $s$ in $s=s$ by~$t$:  | 
|
701  | 
\[ \infer{t=s}{s=t & \infer{s=s}{}} \]
 | 
|
702  | 
||
703  | 
The Isabelle version of the substitution rule looks like this:  | 
|
704  | 
\begin{isabelle}
 | 
|
705  | 
\isasymlbrakk?t\ =\ ?s;\ ?P\ ?s\isasymrbrakk\ \isasymLongrightarrow\ ?P\  | 
|
706  | 
?t  | 
|
| 11417 | 707  | 
\rulenamedx{ssubst}
 | 
| 11077 | 708  | 
\end{isabelle}
 | 
709  | 
Crucially, \isa{?P} is a function 
 | 
|
| 11406 | 710  | 
variable. It can be replaced by a $\lambda$-term  | 
711  | 
with one bound variable, whose occurrences identify the places  | 
|
| 15952 | 712  | 
in which $s$ will be replaced by~$t$.  The proof above requires \isa{?P}
 | 
713  | 
to be replaced by \isa{{\isasymlambda}x.~x=s}; the second premise will then
 | 
|
714  | 
be \isa{s=s} and the conclusion will be \isa{t=s}.
 | 
|
| 11077 | 715  | 
|
| 15952 | 716  | 
The \isa{simp} method also replaces equals by equals, but the substitution
 | 
717  | 
rule gives us more control. Consider this proof:  | 
|
718  | 
\begin{isabelle}
 | 
|
719  | 
\isacommand{lemma}\
 | 
|
720  | 
"\isasymlbrakk x\ =\ f\ x;\ odd(f\ x)\isasymrbrakk\ \isasymLongrightarrow\  | 
|
721  | 
odd\ x"\isanewline  | 
|
722  | 
\isacommand{by}\ (erule\ ssubst)
 | 
|
723  | 
\end{isabelle}
 | 
|
724  | 
%  | 
|
725  | 
The assumption \isa{x\ =\ f\ x}, if used for rewriting, would loop, 
 | 
|
726  | 
replacing \isa{x} by \isa{f x} and then by
 | 
|
727  | 
\isa{f(f x)} and so forth. (Here \isa{simp} 
 | 
|
728  | 
would see the danger and would re-orient the equality, but in more complicated  | 
|
729  | 
cases it can be fooled.) When we apply the substitution rule,  | 
|
730  | 
Isabelle replaces every  | 
|
731  | 
\isa{x} in the subgoal by \isa{f x} just once. It cannot loop.  The
 | 
|
732  | 
resulting subgoal is trivial by assumption, so the \isacommand{by} command
 | 
|
733  | 
proves it implicitly.  | 
|
734  | 
||
735  | 
We are using the \isa{erule} method in a novel way. Hitherto, 
 | 
|
736  | 
the conclusion of the rule was just a variable such as~\isa{?R}, but it may
 | 
|
737  | 
be any term. The conclusion is unified with the subgoal just as  | 
|
738  | 
it would be with the \isa{rule} method. At the same time \isa{erule} looks 
 | 
|
739  | 
for an assumption that matches the rule's first premise, as usual. With  | 
|
740  | 
\isa{ssubst} the effect is to find, use and delete an equality 
 | 
|
741  | 
assumption.  | 
|
742  | 
||
743  | 
The \methdx{subst} method performs individual substitutions. In simple cases,
 | 
|
744  | 
it closely resembles a use of the substitution rule. Suppose a  | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
745  | 
proof has reached this point:  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
746  | 
\begin{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
747  | 
\ 1.\ \isasymlbrakk P\ x\ y\ z;\ Suc\ x\ <\ y\isasymrbrakk \ \isasymLongrightarrow \ f\ z\ =\ x\ *\ y%  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
748  | 
\end{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
749  | 
Now we wish to apply a commutative law:  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
750  | 
\begin{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
751  | 
?m\ *\ ?n\ =\ ?n\ *\ ?m%  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
752  | 
\rulename{mult_commute}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
753  | 
\end{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
754  | 
Isabelle rejects our first attempt:  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
755  | 
\begin{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
756  | 
apply (simp add: mult_commute)  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
757  | 
\end{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
758  | 
The simplifier notices the danger of looping and refuses to apply the  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
759  | 
rule.%  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
760  | 
\footnote{More precisely, it only applies such a rule if the new term is
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
761  | 
smaller under a specified ordering; here, \isa{x\ *\ y}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
762  | 
is already smaller than  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
763  | 
\isa{y\ *\ x}.}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
764  | 
%  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
765  | 
The \isa{subst} method applies \isa{mult_commute} exactly once.  
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
766  | 
\begin{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
767  | 
\isacommand{apply}\ (subst\ mult_commute)\isanewline
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
768  | 
\ 1.\ \isasymlbrakk P\ x\ y\ z;\ Suc\ x\ <\ y\isasymrbrakk \  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
769  | 
\isasymLongrightarrow \ f\ z\ =\ y\ *\ x%  | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
770  | 
\end{isabelle}
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
771  | 
As we wanted, \isa{x\ *\ y} has become \isa{y\ *\ x}.
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
772  | 
|
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
773  | 
\medskip  | 
| 15952 | 774  | 
This use of the \methdx{subst} method has the same effect as the command
 | 
| 11077 | 775  | 
\begin{isabelle}
 | 
| 15952 | 776  | 
\isacommand{apply}\ (rule\ mult_commute [THEN ssubst])
 | 
| 11077 | 777  | 
\end{isabelle}
 | 
| 15952 | 778  | 
The attribute \isa{THEN}, which combines two rules, is described in 
 | 
779  | 
{\S}\ref{sec:THEN} below. The \methdx{subst} method is more powerful than
 | 
|
780  | 
applying the substitution rule. It can perform substitutions in a subgoal's  | 
|
781  | 
assumptions. Moreover, if the subgoal contains more than one occurrence of  | 
|
782  | 
the left-hand side of the equality, the \methdx{subst} method lets us specify which occurrence should be replaced.
 | 
|
| 11077 | 783  | 
|
784  | 
||
785  | 
\subsection{Unification and Its Pitfalls}
 | 
|
786  | 
||
787  | 
Higher-order unification can be tricky. Here is an example, which you may  | 
|
788  | 
want to skip on your first reading:  | 
|
789  | 
\begin{isabelle}
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
790  | 
\isacommand{lemma}\ "\isasymlbrakk x\ =\
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
791  | 
f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\  | 
| 11077 | 792  | 
\isasymLongrightarrow\ triple\ x\ x\ x"\isanewline  | 
793  | 
\isacommand{apply}\ (erule\ ssubst)\isanewline
 | 
|
794  | 
\isacommand{back}\isanewline
 | 
|
795  | 
\isacommand{back}\isanewline
 | 
|
796  | 
\isacommand{back}\isanewline
 | 
|
797  | 
\isacommand{back}\isanewline
 | 
|
798  | 
\isacommand{apply}\ assumption\isanewline
 | 
|
799  | 
\isacommand{done}
 | 
|
800  | 
\end{isabelle}
 | 
|
801  | 
%  | 
|
802  | 
By default, Isabelle tries to substitute for all the  | 
|
803  | 
occurrences.  Applying \isa{erule\ ssubst} yields this subgoal:
 | 
|
804  | 
\begin{isabelle}
 | 
|
805  | 
\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ (f\ x)\ (f\ x)\ (f\ x)  | 
|
806  | 
\end{isabelle}
 | 
|
807  | 
The substitution should have been done in the first two occurrences  | 
|
| 11406 | 808  | 
of~\isa{x} only. Isabelle has gone too far. The \commdx{back}
 | 
809  | 
command allows us to reject this possibility and demand a new one:  | 
|
| 11077 | 810  | 
\begin{isabelle}
 | 
811  | 
\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ x\ (f\ x)\ (f\ x)  | 
|
812  | 
\end{isabelle}
 | 
|
813  | 
%  | 
|
814  | 
Now Isabelle has left the first occurrence of~\isa{x} alone. That is 
 | 
|
| 11406 | 815  | 
promising but it is not the desired combination. So we use \isacommand{back} 
 | 
| 11077 | 816  | 
again:  | 
817  | 
\begin{isabelle}
 | 
|
818  | 
\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ (f\ x)\ x\ (f\ x)  | 
|
819  | 
\end{isabelle}
 | 
|
820  | 
%  | 
|
| 11406 | 821  | 
This also is wrong, so we use \isacommand{back} again: 
 | 
| 11077 | 822  | 
\begin{isabelle}
 | 
823  | 
\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ x\ x\ (f\ x)  | 
|
824  | 
\end{isabelle}
 | 
|
825  | 
%  | 
|
826  | 
And this one is wrong too. Looking carefully at the series  | 
|
827  | 
of alternatives, we see a binary countdown with reversed bits: 111,  | 
|
| 11406 | 828  | 
011, 101, 001.  Invoke \isacommand{back} again: 
 | 
| 11077 | 829  | 
\begin{isabelle}
 | 
830  | 
\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ (f\ x)\ (f\ x)\ x%  | 
|
831  | 
\end{isabelle}
 | 
|
832  | 
At last, we have the right combination! This goal follows by assumption.%  | 
|
833  | 
\index{unification|)}
 | 
|
834  | 
||
| 11406 | 835  | 
\medskip  | 
836  | 
This example shows that unification can do strange things with  | 
|
| 11077 | 837  | 
function variables. We were forced to select the right unifier using the  | 
| 11406 | 838  | 
\isacommand{back} command.  That is all right during exploration, but \isacommand{back}
 | 
| 11077 | 839  | 
should never appear in the final version of a proof. You can eliminate the  | 
| 11406 | 840  | 
need for \isacommand{back} by giving Isabelle less freedom when you apply a rule.
 | 
| 11077 | 841  | 
|
842  | 
One way to constrain the inference is by joining two methods in a  | 
|
843  | 
\isacommand{apply} command. Isabelle  applies the first method and then the
 | 
|
844  | 
second. If the second method fails then Isabelle automatically backtracks.  | 
|
845  | 
This process continues until the first method produces an output that the  | 
|
846  | 
second method can use. We get a one-line proof of our example:  | 
|
847  | 
\begin{isabelle}
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
848  | 
\isacommand{lemma}\ "\isasymlbrakk x\ =\ f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\
 | 
| 11077 | 849  | 
\isasymLongrightarrow\ triple\ x\ x\ x"\isanewline  | 
850  | 
\isacommand{apply}\ (erule\ ssubst,\ assumption)\isanewline
 | 
|
851  | 
\isacommand{done}
 | 
|
852  | 
\end{isabelle}
 | 
|
853  | 
||
854  | 
\noindent  | 
|
855  | 
The \isacommand{by} command works too, since it backtracks when
 | 
|
856  | 
proving subgoals by assumption:  | 
|
857  | 
\begin{isabelle}
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
858  | 
\isacommand{lemma}\ "\isasymlbrakk x\ =\ f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\
 | 
| 11077 | 859  | 
\isasymLongrightarrow\ triple\ x\ x\ x"\isanewline  | 
860  | 
\isacommand{by}\ (erule\ ssubst)
 | 
|
861  | 
\end{isabelle}
 | 
|
862  | 
||
863  | 
||
864  | 
The most general way to constrain unification is  | 
|
865  | 
by instantiating variables in the rule.  The method \isa{rule_tac} is
 | 
|
866  | 
similar to \isa{rule}, but it
 | 
|
867  | 
makes some of the rule's variables denote specified terms.  | 
|
868  | 
Also available are {\isa{drule_tac}}  and \isa{erule_tac}.  Here we need
 | 
|
869  | 
\isa{erule_tac} since above we used \isa{erule}.
 | 
|
870  | 
\begin{isabelle}
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
871  | 
\isacommand{lemma}\ "\isasymlbrakk x\ =\ f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\ \isasymLongrightarrow\ triple\ x\ x\ x"\isanewline
 | 
| 
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
872  | 
\isacommand{by}\ (erule_tac\ P = "\isasymlambda u.\ triple\ u\ u\ x"\ 
 | 
| 11077 | 873  | 
\isakeyword{in}\ ssubst)
 | 
874  | 
\end{isabelle}
 | 
|
875  | 
%  | 
|
876  | 
To specify a desired substitution  | 
|
877  | 
requires instantiating the variable \isa{?P} with a $\lambda$-expression. 
 | 
|
878  | 
The bound variable occurrences in \isa{{\isasymlambda}u.\ P\ u\
 | 
|
879  | 
u\ x} indicate that the first two arguments have to be substituted, leaving  | 
|
880  | 
the third unchanged. With this instantiation, backtracking is neither necessary  | 
|
881  | 
nor possible.  | 
|
882  | 
||
| 11406 | 883  | 
An alternative to \isa{rule_tac} is to use \isa{rule} with a theorem
 | 
884  | 
modified using~\isa{of}, described in
 | 
|
| 12540 | 885  | 
{\S}\ref{sec:forward} below.   But \isa{rule_tac}, unlike \isa{of}, can 
 | 
886  | 
express instantiations that refer to  | 
|
| 11077 | 887  | 
\isasymAnd-bound variables in the current subgoal.%  | 
888  | 
\index{substitution|)}
 | 
|
| 10295 | 889  | 
|
890  | 
||
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
891  | 
\section{Quantifiers}
 | 
| 10295 | 892  | 
|
| 11411 | 893  | 
\index{quantifiers!universal|(}%
 | 
| 11077 | 894  | 
Quantifiers require formalizing syntactic substitution and the notion of  | 
| 11406 | 895  | 
arbitrary value. Consider the universal quantifier. In a logic  | 
| 11077 | 896  | 
book, its introduction rule looks like this:  | 
| 10295 | 897  | 
\[ \infer{\forall x.\,P}{P} \]
 | 
898  | 
Typically, a proviso written in English says that $x$ must not  | 
|
899  | 
occur in the assumptions. This proviso guarantees that $x$ can be regarded as  | 
|
900  | 
arbitrary, since it has not been assumed to satisfy any special conditions.  | 
|
901  | 
Isabelle's underlying formalism, called the  | 
|
| 11406 | 902  | 
\bfindex{meta-logic}, eliminates the  need for English.  It provides its own
 | 
| 27167 | 903  | 
universal quantifier (\isasymAnd) to express the notion of an arbitrary value.  | 
904  | 
We have already seen another operator of the meta-logic, namely  | 
|
905  | 
\isa\isasymLongrightarrow, which expresses inference rules and the treatment  | 
|
906  | 
of assumptions. The only other operator in the meta-logic is \isa\isasymequiv,  | 
|
907  | 
which can be used to define constants.  | 
|
| 10295 | 908  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
909  | 
\subsection{The Universal Introduction Rule}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
910  | 
|
| 10295 | 911  | 
Returning to the universal quantifier, we find that having a similar quantifier  | 
912  | 
as part of the meta-logic makes the introduction rule trivial to express:  | 
|
913  | 
\begin{isabelle}
 | 
|
| 11417 | 914  | 
(\isasymAnd x.\ ?P\ x)\ \isasymLongrightarrow\ {\isasymforall}x.\ ?P\ x\rulenamedx{allI}
 | 
| 10295 | 915  | 
\end{isabelle}
 | 
916  | 
||
917  | 
||
918  | 
The following trivial proof demonstrates how the universal introduction  | 
|
919  | 
rule works.  | 
|
920  | 
\begin{isabelle}
 | 
|
921  | 
\isacommand{lemma}\ "{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ x"\isanewline
 | 
|
922  | 
\isacommand{apply}\ (rule\ allI)\isanewline
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
923  | 
\isacommand{by}\ (rule\ impI)
 | 
| 10295 | 924  | 
\end{isabelle}
 | 
925  | 
The first step invokes the rule by applying the method \isa{rule allI}. 
 | 
|
926  | 
\begin{isabelle}
 | 
|
| 10596 | 927  | 
\ 1.\ \isasymAnd x.\ P\ x\ \isasymlongrightarrow\ P\ x  | 
| 10295 | 928  | 
\end{isabelle}
 | 
929  | 
Note that the resulting proof state has a bound variable,  | 
|
| 11077 | 930  | 
namely~\isa{x}.  The rule has replaced the universal quantifier of
 | 
| 10295 | 931  | 
higher-order logic by Isabelle's meta-level quantifier. Our goal is to  | 
932  | 
prove  | 
|
933  | 
\isa{P\ x\ \isasymlongrightarrow\ P\ x} for arbitrary~\isa{x}; it is 
 | 
|
934  | 
an implication, so we apply the corresponding introduction rule (\isa{impI}). 
 | 
|
935  | 
\begin{isabelle}
 | 
|
| 10596 | 936  | 
\ 1.\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow\ P\ x  | 
| 10295 | 937  | 
\end{isabelle}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
938  | 
This last subgoal is implicitly proved by assumption.  | 
| 10295 | 939  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
940  | 
\subsection{The Universal Elimination Rule}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
941  | 
|
| 10295 | 942  | 
Now consider universal elimination. In a logic text,  | 
943  | 
the rule looks like this:  | 
|
944  | 
\[ \infer{P[t/x]}{\forall x.\,P} \]
 | 
|
945  | 
The conclusion is $P$ with $t$ substituted for the variable~$x$.  | 
|
946  | 
Isabelle expresses substitution using a function variable:  | 
|
947  | 
\begin{isabelle}
 | 
|
| 11417 | 948  | 
{\isasymforall}x.\ ?P\ x\ \isasymLongrightarrow\ ?P\ ?x\rulenamedx{spec}
 | 
| 10295 | 949  | 
\end{isabelle}
 | 
950  | 
This destruction rule takes a  | 
|
951  | 
universally quantified formula and removes the quantifier, replacing  | 
|
| 11077 | 952  | 
the bound variable \isa{x} by the schematic variable \isa{?x}.  Recall that a
 | 
| 10295 | 953  | 
schematic variable starts with a question mark and acts as a  | 
| 11077 | 954  | 
placeholder: it can be replaced by any term.  | 
| 10295 | 955  | 
|
| 11077 | 956  | 
The universal elimination rule is also  | 
957  | 
available in the standard elimination format.  Like \isa{conjE}, it never
 | 
|
958  | 
appears in logic books:  | 
|
959  | 
\begin{isabelle}
 | 
|
960  | 
\isasymlbrakk \isasymforall x.\ ?P\ x;\ ?P\ ?x\ \isasymLongrightarrow \ ?R\isasymrbrakk \ \isasymLongrightarrow \ ?R%  | 
|
| 11417 | 961  | 
\rulenamedx{allE}
 | 
| 11077 | 962  | 
\end{isabelle}
 | 
963  | 
The methods \isa{drule~spec} and \isa{erule~allE} do precisely the
 | 
|
964  | 
same inference.  | 
|
965  | 
||
966  | 
To see how $\forall$-elimination works, let us derive a rule about reducing  | 
|
| 10295 | 967  | 
the scope of a universal quantifier. In mathematical notation we write  | 
968  | 
\[ \infer{P\imp\forall x.\,Q}{\forall x.\,P\imp Q} \]
 | 
|
| 10978 | 969  | 
with the proviso ``$x$ not free in~$P$.'' Isabelle's treatment of  | 
| 10295 | 970  | 
substitution makes the proviso unnecessary. The conclusion is expressed as  | 
971  | 
\isa{P\
 | 
|
972  | 
\isasymlongrightarrow\ ({\isasymforall}x.\ Q\ x)}. No substitution for the
 | 
|
973  | 
variable \isa{P} can introduce a dependence upon~\isa{x}: that would be a
 | 
|
| 11077 | 974  | 
bound variable capture. Let us walk through the proof.  | 
| 10295 | 975  | 
\begin{isabelle}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
976  | 
\isacommand{lemma}\ "(\isasymforall x.\ P\ \isasymlongrightarrow \ Q\ x)\
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
977  | 
\isasymLongrightarrow \ P\ \isasymlongrightarrow \ (\isasymforall x.\ Q\  | 
| 11077 | 978  | 
x)"  | 
| 10295 | 979  | 
\end{isabelle}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
980  | 
First we apply implies introduction (\isa{impI}), 
 | 
| 10295 | 981  | 
which moves the \isa{P} from the conclusion to the assumptions. Then 
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
982  | 
we apply universal introduction (\isa{allI}).  
 | 
| 10295 | 983  | 
\begin{isabelle}
 | 
| 11077 | 984  | 
\isacommand{apply}\ (rule\ impI,\ rule\ allI)\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
985  | 
\ 1.\ \isasymAnd x.\ \isasymlbrakk{\isasymforall}x.\ P\ \isasymlongrightarrow\ Q\
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
986  | 
x;\ P\isasymrbrakk\ \isasymLongrightarrow\ Q\ x  | 
| 10295 | 987  | 
\end{isabelle}
 | 
988  | 
As before, it replaces the HOL  | 
|
989  | 
quantifier by a meta-level quantifier, producing a subgoal that  | 
|
| 11077 | 990  | 
binds the variable~\isa{x}.  The leading bound variables
 | 
| 10295 | 991  | 
(here \isa{x}) and the assumptions (here \isa{{\isasymforall}x.\ P\
 | 
| 11406 | 992  | 
\isasymlongrightarrow\ Q\ x} and \isa{P}) form the \textbf{context} for the
 | 
| 11077 | 993  | 
conclusion, here \isa{Q\ x}.  Subgoals inherit the context,
 | 
994  | 
although assumptions can be added or deleted (as we saw  | 
|
995  | 
earlier), while rules such as \isa{allI} add bound variables.
 | 
|
| 10295 | 996  | 
|
997  | 
Now, to reason from the universally quantified  | 
|
| 10967 | 998  | 
assumption, we apply the elimination rule using the \isa{drule} 
 | 
| 10295 | 999  | 
method.  This rule is called \isa{spec} because it specializes a universal formula
 | 
1000  | 
to a particular term.  | 
|
1001  | 
\begin{isabelle}
 | 
|
| 11077 | 1002  | 
\isacommand{apply}\ (drule\ spec)\isanewline
 | 
| 10596 | 1003  | 
\ 1.\ \isasymAnd x.\ \isasymlbrakk P;\ P\ \isasymlongrightarrow\ Q\ (?x2\  | 
1004  | 
x)\isasymrbrakk\ \isasymLongrightarrow\ Q\ x  | 
|
| 10295 | 1005  | 
\end{isabelle}
 | 
1006  | 
Observe how the context has changed. The quantified formula is gone,  | 
|
| 11406 | 1007  | 
replaced by a new assumption derived from its body. We have  | 
1008  | 
removed the quantifier and replaced the bound variable  | 
|
1009  | 
by the curious term  | 
|
1010  | 
\isa{?x2~x}.  This term is a placeholder: it may become any term that can be
 | 
|
1011  | 
built from~\isa{x}.  (Formally, \isa{?x2} is an unknown of function type, applied
 | 
|
1012  | 
to the argument~\isa{x}.)  This new assumption is an implication, so we can  use
 | 
|
1013  | 
\emph{modus ponens} on it, which concludes the proof. 
 | 
|
| 11077 | 1014  | 
\begin{isabelle}
 | 
1015  | 
\isacommand{by}\ (drule\ mp)
 | 
|
1016  | 
\end{isabelle}
 | 
|
1017  | 
Let us take a closer look at this last step.  \emph{Modus ponens} yields
 | 
|
1018  | 
two subgoals: one where we prove the antecedent (in this case \isa{P}) and
 | 
|
1019  | 
one where we may assume the consequent. Both of these subgoals are proved  | 
|
1020  | 
by the  | 
|
1021  | 
\isa{assumption} method, which is implicit in the
 | 
|
1022  | 
\isacommand{by} command.  Replacing the \isacommand{by} command by 
 | 
|
1023  | 
\isa{\isacommand{apply} (drule\ mp, assumption)} would have left one last
 | 
|
1024  | 
subgoal:  | 
|
| 10295 | 1025  | 
\begin{isabelle}
 | 
| 10596 | 1026  | 
\ 1.\ \isasymAnd x.\ \isasymlbrakk P;\ Q\ (?x2\ x)\isasymrbrakk\  | 
| 10295 | 1027  | 
\isasymLongrightarrow\ Q\ x  | 
1028  | 
\end{isabelle}
 | 
|
1029  | 
The consequent is \isa{Q} applied to that placeholder.  It may be replaced by any
 | 
|
| 11077 | 1030  | 
term built from~\isa{x}, and here 
 | 
1031  | 
it should simply be~\isa{x}.  The assumption need not
 | 
|
1032  | 
be identical to the conclusion, provided the two formulas are unifiable.%  | 
|
1033  | 
\index{quantifiers!universal|)}  
 | 
|
| 10295 | 1034  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1035  | 
|
| 11234 | 1036  | 
\subsection{The Existential Quantifier}
 | 
1037  | 
||
1038  | 
\index{quantifiers!existential|(}%
 | 
|
1039  | 
The concepts just presented also apply  | 
|
1040  | 
to the existential quantifier, whose introduction rule looks like this in  | 
|
1041  | 
Isabelle:  | 
|
1042  | 
\begin{isabelle}
 | 
|
| 11417 | 1043  | 
?P\ ?x\ \isasymLongrightarrow\ {\isasymexists}x.\ ?P\ x\rulenamedx{exI}
 | 
| 11234 | 1044  | 
\end{isabelle}
 | 
1045  | 
If we can exhibit some $x$ such that $P(x)$ is true, then $\exists x.  | 
|
1046  | 
P(x)$ is also true. It is a dual of the universal elimination rule, and  | 
|
1047  | 
logic texts present it using the same notation for substitution.  | 
|
1048  | 
||
1049  | 
The existential  | 
|
1050  | 
elimination rule looks like this  | 
|
1051  | 
in a logic text:  | 
|
1052  | 
\[ \infer{Q}{\exists x.\,P & \infer*{Q}{[P]}} \]
 | 
|
1053  | 
%  | 
|
1054  | 
It looks like this in Isabelle:  | 
|
1055  | 
\begin{isabelle}
 | 
|
| 11417 | 1056  | 
\isasymlbrakk{\isasymexists}x.\ ?P\ x;\ \isasymAnd x.\ ?P\ x\ \isasymLongrightarrow\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?Q\rulenamedx{exE}
 | 
| 11234 | 1057  | 
\end{isabelle}
 | 
1058  | 
%  | 
|
1059  | 
Given an existentially quantified theorem and some  | 
|
1060  | 
formula $Q$ to prove, it creates a new assumption by removing the quantifier. As with  | 
|
1061  | 
the universal introduction rule, the textbook version imposes a proviso on the  | 
|
| 11406 | 1062  | 
quantified variable, which Isabelle expresses using its meta-logic. It is  | 
| 11234 | 1063  | 
enough to have a universal quantifier in the meta-logic; we do not need an existential  | 
1064  | 
quantifier to be built in as well.  | 
|
1065  | 
||
1066  | 
||
1067  | 
\begin{exercise}
 | 
|
1068  | 
Prove the lemma  | 
|
1069  | 
\[ \exists x.\, P\conj Q(x)\Imp P\conj(\exists x.\, Q(x)). \]  | 
|
1070  | 
\emph{Hint}: the proof is similar 
 | 
|
1071  | 
to the one just above for the universal quantifier.  | 
|
1072  | 
\end{exercise}
 | 
|
| 11411 | 1073  | 
\index{quantifiers!existential|)}
 | 
| 11234 | 1074  | 
|
1075  | 
||
| 34054 | 1076  | 
\subsection{Renaming a Bound Variable: {\tt\slshape rename_tac}}
 | 
| 10967 | 1077  | 
|
| 11406 | 1078  | 
\index{assumptions!renaming|(}\index{*rename_tac (method)|(}%
 | 
| 11077 | 1079  | 
When you apply a rule such as \isa{allI}, the quantified variable
 | 
1080  | 
becomes a new bound variable of the new subgoal. Isabelle tries to avoid  | 
|
1081  | 
changing its name, but sometimes it has to choose a new name in order to  | 
|
| 11234 | 1082  | 
avoid a clash. The result may not be ideal:  | 
| 10967 | 1083  | 
\begin{isabelle}
 | 
1084  | 
\isacommand{lemma}\ "x\ <\ y\ \isasymLongrightarrow \ \isasymforall x\ y.\ P\ x\
 | 
|
1085  | 
(f\ y)"\isanewline  | 
|
| 12408 | 1086  | 
\isacommand{apply}\ (intro allI)\isanewline
 | 
| 10967 | 1087  | 
\ 1.\ \isasymAnd xa\ ya.\ x\ <\ y\ \isasymLongrightarrow \ P\ xa\ (f\ ya)  | 
1088  | 
\end{isabelle}
 | 
|
1089  | 
%  | 
|
1090  | 
The names \isa{x} and \isa{y} were already in use, so the new bound variables are
 | 
|
1091  | 
called \isa{xa} and~\isa{ya}.  You can rename them by invoking \isa{rename_tac}:
 | 
|
1092  | 
||
1093  | 
\begin{isabelle}
 | 
|
1094  | 
\isacommand{apply}\ (rename_tac\ v\ w)\isanewline
 | 
|
1095  | 
\ 1.\ \isasymAnd v\ w.\ x\ <\ y\ \isasymLongrightarrow \ P\ v\ (f\ w)  | 
|
1096  | 
\end{isabelle}
 | 
|
| 11406 | 1097  | 
Recall that \isa{rule_tac}\index{*rule_tac (method)!and renaming} 
 | 
1098  | 
instantiates a  | 
|
| 10967 | 1099  | 
theorem with specified terms. These terms may involve the goal's bound  | 
1100  | 
variables, but beware of referring to variables  | 
|
1101  | 
like~\isa{xa}.  A future change to your theories could change the set of names
 | 
|
1102  | 
produced at top level, so that \isa{xa} changes to~\isa{xb} or reverts to~\isa{x}.
 | 
|
1103  | 
It is safer to rename automatically-generated variables before mentioning them.  | 
|
1104  | 
||
1105  | 
If the subgoal has more bound variables than there are names given to  | 
|
| 11077 | 1106  | 
\isa{rename_tac}, the rightmost ones are renamed.%
 | 
| 11406 | 1107  | 
\index{assumptions!renaming|)}\index{*rename_tac (method)|)}
 | 
| 10967 | 1108  | 
|
1109  | 
||
1110  | 
\subsection{Reusing an Assumption: {\tt\slshape frule}}
 | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
1111  | 
\label{sec:frule}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1112  | 
|
| 11406 | 1113  | 
\index{assumptions!reusing|(}\index{*frule (method)|(}%
 | 
| 10295 | 1114  | 
Note that \isa{drule spec} removes the universal quantifier and --- as
 | 
1115  | 
usual with elimination rules --- discards the original formula. Sometimes, a  | 
|
1116  | 
universal formula has to be kept so that it can be used again. Then we use a new  | 
|
1117  | 
method: \isa{frule}.  It acts like \isa{drule} but copies rather than replaces
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1118  | 
the selected assumption.  The \isa{f} is for \emph{forward}.
 | 
| 10295 | 1119  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1120  | 
In this example, going from \isa{P\ a} to \isa{P(h(h~a))}
 | 
| 11406 | 1121  | 
requires two uses of the quantified assumption, one for each~\isa{h}
 | 
1122  | 
in~\isa{h(h~a)}.
 | 
|
| 10295 | 1123  | 
\begin{isabelle}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1124  | 
\isacommand{lemma}\ "\isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\ x);
 | 
| 11077 | 1125  | 
\ P\ a\isasymrbrakk\ \isasymLongrightarrow\ P(h\ (h\ a))"  | 
| 10295 | 1126  | 
\end{isabelle}
 | 
1127  | 
%  | 
|
| 11077 | 1128  | 
Examine the subgoal left by \isa{frule}:
 | 
| 10295 | 1129  | 
\begin{isabelle}
 | 
| 11077 | 1130  | 
\isacommand{apply}\ (frule\ spec)\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1131  | 
\ 1.\ \isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\ x);\ P\ a;\ P\ ?x\ \isasymlongrightarrow\ P\ (h\ ?x)\isasymrbrakk\ \isasymLongrightarrow\ P\ (h\ (h\ a))
 | 
| 10295 | 1132  | 
\end{isabelle}
 | 
| 11077 | 1133  | 
It is what \isa{drule} would have left except that the quantified
 | 
1134  | 
assumption is still present.  Next we apply \isa{mp} to the
 | 
|
1135  | 
implication and the assumption~\isa{P\ a}:
 | 
|
| 10295 | 1136  | 
\begin{isabelle}
 | 
| 11077 | 1137  | 
\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1138  | 
\ 1.\ \isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\ x);\ P\ a;\ P\ (h\ a)\isasymrbrakk\ \isasymLongrightarrow\ P\ (h\ (h\ a))
 | 
| 10295 | 1139  | 
\end{isabelle}
 | 
1140  | 
%  | 
|
| 11077 | 1141  | 
We have created the assumption \isa{P(h\ a)}, which is progress.  To
 | 
1142  | 
continue the proof, we apply \isa{spec} again.  We shall not need it
 | 
|
1143  | 
again, so we can use  | 
|
1144  | 
\isa{drule}.
 | 
|
1145  | 
\begin{isabelle}
 | 
|
1146  | 
\isacommand{apply}\ (drule\ spec)\isanewline
 | 
|
1147  | 
\ 1.\ \isasymlbrakk P\ a;\ P\ (h\ a);\ P\ ?x2\  | 
|
1148  | 
\isasymlongrightarrow \ P\ (h\ ?x2)\isasymrbrakk \ \isasymLongrightarrow \  | 
|
1149  | 
P\ (h\ (h\ a))  | 
|
1150  | 
\end{isabelle}
 | 
|
1151  | 
%  | 
|
1152  | 
The new assumption bridges the gap between \isa{P(h\ a)} and \isa{P(h(h\ a))}.
 | 
|
1153  | 
\begin{isabelle}
 | 
|
1154  | 
\isacommand{by}\ (drule\ mp)
 | 
|
1155  | 
\end{isabelle}
 | 
|
| 10854 | 1156  | 
|
1157  | 
\medskip  | 
|
| 11077 | 1158  | 
\emph{A final remark}.  Replacing this \isacommand{by} command with
 | 
| 10295 | 1159  | 
\begin{isabelle}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1160  | 
\isacommand{apply}\ (drule\ mp,\ assumption)
 | 
| 10295 | 1161  | 
\end{isabelle}
 | 
| 11077 | 1162  | 
would not work: it would add a second copy of \isa{P(h~a)} instead
 | 
| 10854 | 1163  | 
of the desired assumption, \isa{P(h(h~a))}.  The \isacommand{by}
 | 
1164  | 
command forces Isabelle to backtrack until it finds the correct one.  | 
|
1165  | 
Alternatively, we could have used the \isacommand{apply} command and bundled the
 | 
|
| 11234 | 1166  | 
\isa{drule mp} with \emph{two} calls of \isa{assumption}.  Or, of course,
 | 
1167  | 
we could have given the entire proof to \isa{auto}.%
 | 
|
| 11406 | 1168  | 
\index{assumptions!reusing|)}\index{*frule (method)|)}
 | 
| 10295 | 1169  | 
|
1170  | 
||
| 11234 | 1171  | 
|
1172  | 
\subsection{Instantiating a Quantifier Explicitly}
 | 
|
1173  | 
\index{quantifiers!instantiating}
 | 
|
| 10295 | 1174  | 
|
| 11234 | 1175  | 
We can prove a theorem of the form $\exists x.\,P\, x$ by exhibiting a  | 
1176  | 
suitable term~$t$ such that $P\,t$ is true. Dually, we can use an  | 
|
| 11406 | 1177  | 
assumption of the form $\forall x.\,P\, x$ to generate a new assumption $P\,t$ for  | 
1178  | 
a suitable term~$t$. In many cases,  | 
|
| 11234 | 1179  | 
Isabelle makes the correct choice automatically, constructing the term by  | 
1180  | 
unification. In other cases, the required term is not obvious and we must  | 
|
1181  | 
specify it ourselves.  Suitable methods are \isa{rule_tac}, \isa{drule_tac}
 | 
|
1182  | 
and \isa{erule_tac}.
 | 
|
1183  | 
||
| 11428 | 1184  | 
We have seen (just above, {\S}\ref{sec:frule}) a proof of this lemma:
 | 
| 10295 | 1185  | 
\begin{isabelle}
 | 
| 11234 | 1186  | 
\isacommand{lemma}\ "\isasymlbrakk \isasymforall x.\ P\ x\
 | 
1187  | 
\isasymlongrightarrow \ P\ (h\ x);\ P\ a\isasymrbrakk \  | 
|
1188  | 
\isasymLongrightarrow \ P(h\ (h\ a))"  | 
|
| 10295 | 1189  | 
\end{isabelle}
 | 
| 11234 | 1190  | 
We had reached this subgoal:  | 
| 10295 | 1191  | 
\begin{isabelle}
 | 
| 11234 | 1192  | 
\ 1.\ \isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\
 | 
1193  | 
x);\ P\ a;\ P\ (h\ a)\isasymrbrakk\ \isasymLongrightarrow\ P\ (h\ (h\ a))  | 
|
| 10295 | 1194  | 
\end{isabelle}
 | 
1195  | 
%  | 
|
| 11234 | 1196  | 
The proof requires instantiating the quantified assumption with the  | 
1197  | 
term~\isa{h~a}.
 | 
|
1198  | 
\begin{isabelle}
 | 
|
1199  | 
\isacommand{apply}\ (drule_tac\ x\ =\ "h\ a"\ \isakeyword{in}\
 | 
|
1200  | 
spec)\isanewline  | 
|
1201  | 
\ 1.\ \isasymlbrakk P\ a;\ P\ (h\ a);\ P\ (h\ a)\ \isasymlongrightarrow \  | 
|
1202  | 
P\ (h\ (h\ a))\isasymrbrakk \ \isasymLongrightarrow \ P\ (h\ (h\ a))  | 
|
1203  | 
\end{isabelle}
 | 
|
1204  | 
We have forced the desired instantiation.  | 
|
1205  | 
||
1206  | 
\medskip  | 
|
1207  | 
Existential formulas can be instantiated too. The next example uses the  | 
|
| 11417 | 1208  | 
\textbf{divides} relation\index{divides relation}
 | 
| 11406 | 1209  | 
of number theory:  | 
| 11234 | 1210  | 
\begin{isabelle}
 | 
1211  | 
?m\ dvd\ ?n\ \isasymequiv\ {\isasymexists}k.\ ?n\ =\ ?m\ *\ k
 | 
|
1212  | 
\rulename{dvd_def}
 | 
|
1213  | 
\end{isabelle}
 | 
|
| 10295 | 1214  | 
|
| 11234 | 1215  | 
Let us prove that multiplication of natural numbers is monotone with  | 
1216  | 
respect to the divides relation:  | 
|
1217  | 
\begin{isabelle}
 | 
|
1218  | 
\isacommand{lemma}\ mult_dvd_mono:\ "{\isasymlbrakk}i\ dvd\ m;\ j\ dvd\
 | 
|
1219  | 
n\isasymrbrakk\ \isasymLongrightarrow\ i*j\ dvd\ (m*n\ ::\ nat)"\isanewline  | 
|
1220  | 
\isacommand{apply}\ (simp\ add:\ dvd_def)
 | 
|
1221  | 
\end{isabelle}
 | 
|
1222  | 
%  | 
|
| 11406 | 1223  | 
Unfolding the definition of divides has left this subgoal:  | 
| 11234 | 1224  | 
\begin{isabelle}
 | 
1225  | 
\ 1.\ \isasymlbrakk \isasymexists k.\ m\ =\ i\ *\ k;\ \isasymexists k.\ n\  | 
|
1226  | 
=\ j\ *\ k\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\ m\ *\  | 
|
| 11406 | 1227  | 
n\ =\ i\ *\ j\ *\ k  | 
1228  | 
\end{isabelle}
 | 
|
1229  | 
%  | 
|
1230  | 
Next, we eliminate the two existential quantifiers in the assumptions:  | 
|
1231  | 
\begin{isabelle}
 | 
|
| 11234 | 1232  | 
\isacommand{apply}\ (erule\ exE)\isanewline
 | 
1233  | 
\ 1.\ \isasymAnd k.\ \isasymlbrakk \isasymexists k.\ n\ =\ j\ *\ k;\ m\ =\  | 
|
1234  | 
i\ *\ k\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\ m\ *\ n\  | 
|
1235  | 
=\ i\ *\ j\ *\ k%  | 
|
1236  | 
\isanewline  | 
|
1237  | 
\isacommand{apply}\ (erule\ exE)
 | 
|
| 11406 | 1238  | 
\isanewline  | 
| 11234 | 1239  | 
\ 1.\ \isasymAnd k\ ka.\ \isasymlbrakk m\ =\ i\ *\ k;\ n\ =\ j\ *\  | 
1240  | 
ka\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\ m\ *\ n\ =\ i\  | 
|
1241  | 
*\ j\ *\ k  | 
|
1242  | 
\end{isabelle}
 | 
|
1243  | 
%  | 
|
| 11406 | 1244  | 
The term needed to instantiate the remaining quantifier is~\isa{k*ka}.  But
 | 
1245  | 
\isa{ka} is an automatically-generated name.  As noted above, references to
 | 
|
1246  | 
such variable names makes a proof less resilient to future changes. So,  | 
|
1247  | 
first we rename the most recent variable to~\isa{l}:
 | 
|
| 11234 | 1248  | 
\begin{isabelle}
 | 
| 11406 | 1249  | 
\isacommand{apply}\ (rename_tac\ l)\isanewline
 | 
1250  | 
\ 1.\ \isasymAnd k\ l.\ \isasymlbrakk m\ =\ i\ *\ k;\ n\ =\ j\ *\ l\isasymrbrakk \  | 
|
1251  | 
\isasymLongrightarrow \ \isasymexists k.\ m\ *\ n\ =\ i\ *\ j\ *\ k%  | 
|
1252  | 
\end{isabelle}
 | 
|
1253  | 
||
1254  | 
We instantiate the quantifier with~\isa{k*l}:
 | 
|
1255  | 
\begin{isabelle}
 | 
|
1256  | 
\isacommand{apply}\ (rule_tac\ x="k*l"\ \isakeyword{in}\ exI)\ \isanewline
 | 
|
| 11234 | 1257  | 
\ 1.\ \isasymAnd k\ ka.\ \isasymlbrakk m\ =\ i\ *\ k;\ n\ =\ j\ *\  | 
1258  | 
ka\isasymrbrakk \ \isasymLongrightarrow \ m\ *\ n\ =\ i\  | 
|
1259  | 
*\ j\ *\ (k\ *\ ka)  | 
|
1260  | 
\end{isabelle}
 | 
|
1261  | 
%  | 
|
1262  | 
The rest is automatic, by arithmetic.  | 
|
1263  | 
\begin{isabelle}
 | 
|
1264  | 
\isacommand{apply}\ simp\isanewline
 | 
|
1265  | 
\isacommand{done}\isanewline
 | 
|
1266  | 
\end{isabelle}
 | 
|
1267  | 
||
1268  | 
||
| 10295 | 1269  | 
|
| 11458 | 1270  | 
\section{Description Operators}
 | 
| 10971 | 1271  | 
\label{sec:SOME}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1272  | 
|
| 11458 | 1273  | 
\index{description operators|(}%
 | 
1274  | 
HOL provides two description operators.  | 
|
1275  | 
A \textbf{definite description} formalizes the word ``the,'' as in
 | 
|
1276  | 
``the greatest divisior of~$n$.''  | 
|
1277  | 
It returns an arbitrary value unless the formula has a unique solution.  | 
|
1278  | 
An \textbf{indefinite description} formalizes the word ``some,'' as in
 | 
|
| 12815 | 1279  | 
``some member of~$S$.'' It differs from a definite description in not  | 
| 11458 | 1280  | 
requiring the solution to be unique: it uses the axiom of choice to pick any  | 
1281  | 
solution.  | 
|
| 11077 | 1282  | 
|
1283  | 
\begin{warn}
 | 
|
| 11458 | 1284  | 
Description operators can be hard to reason about. Novices  | 
1285  | 
should try to avoid them. Fortunately, descriptions are seldom required.  | 
|
| 11077 | 1286  | 
\end{warn}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1287  | 
|
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1288  | 
\subsection{Definite Descriptions}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1289  | 
|
| 11077 | 1290  | 
\index{descriptions!definite}%
 | 
| 11458 | 1291  | 
A definite description is traditionally written $\iota x. P(x)$. It denotes  | 
1292  | 
the $x$ such that $P(x)$ is true, provided there exists a unique such~$x$;  | 
|
1293  | 
otherwise, it returns an arbitrary value of the expected type.  | 
|
| 12540 | 1294  | 
Isabelle uses \sdx{THE} for the Greek letter~$\iota$.  
 | 
1295  | 
||
1296  | 
%(The traditional notation could be provided, but it is not legible on screen.)  | 
|
| 11458 | 1297  | 
|
1298  | 
We reason using this rule, where \isa{a} is the unique solution:
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1299  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1300  | 
\isasymlbrakk P\ a;\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow \ x\ =\ a\isasymrbrakk \  | 
| 11458 | 1301  | 
\isasymLongrightarrow \ (THE\ x.\ P\ x)\ =\ a%  | 
1302  | 
\rulenamedx{the_equality}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1303  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1304  | 
For instance, we can define the  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1305  | 
cardinality of a finite set~$A$ to be that  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1306  | 
$n$ such that $A$ is in one-to-one correspondence with $\{1,\ldots,n\}$.  We can then
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1307  | 
prove that the cardinality of the empty set is zero (since $n=0$ satisfies the  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1308  | 
description) and proceed to prove other facts.  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1309  | 
|
| 11406 | 1310  | 
A more challenging example illustrates how Isabelle/HOL defines the least number  | 
1311  | 
operator, which denotes the least \isa{x} satisfying~\isa{P}:%
 | 
|
| 11428 | 1312  | 
\index{least number operator|see{\protect\isa{LEAST}}}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1313  | 
\begin{isabelle}
 | 
| 11458 | 1314  | 
(LEAST\ x.\ P\ x)\ = (THE\ x.\ P\ x\ \isasymand \ (\isasymforall y.\  | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
1315  | 
P\ y\ \isasymlongrightarrow \ x\ \isasymle \ y))  | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1316  | 
\end{isabelle}
 | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
1317  | 
%  | 
| 11458 | 1318  | 
Let us prove the analogue of \isa{the_equality} for \sdx{LEAST}\@.
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1319  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1320  | 
\isacommand{theorem}\ Least_equality:\isanewline
 | 
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
1321  | 
\ \ \ \ \ "\isasymlbrakk P\ (k::nat);\ \ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\ \isasymle \ x\isasymrbrakk \ \isasymLongrightarrow \ (LEAST\ x.\ P\ x)\ =\ k"\isanewline  | 
| 11458 | 1322  | 
\isacommand{apply}\ (simp\ add:\ Least_def)\isanewline
 | 
1323  | 
\isanewline  | 
|
1324  | 
\ 1.\ \isasymlbrakk P\ k;\ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\ \isasymle \ x\isasymrbrakk \isanewline  | 
|
1325  | 
\isaindent{\ 1.\ }\isasymLongrightarrow \ (THE\ x.\ P\ x\ \isasymand \ (\isasymforall y.\ P\ y\ \isasymlongrightarrow \ x\ \isasymle \ y))\ =\ k%
 | 
|
1326  | 
\end{isabelle}
 | 
|
1327  | 
The first step has merely unfolded the definition.  | 
|
1328  | 
\begin{isabelle}
 | 
|
1329  | 
\isacommand{apply}\ (rule\ the_equality)\isanewline
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1330  | 
\isanewline  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1331  | 
\ 1.\ \isasymlbrakk P\ k;\ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1332  | 
\isasymle \ x\isasymrbrakk \ \isasymLongrightarrow \ P\ k\ \isasymand \  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1333  | 
(\isasymforall y.\ P\ y\ \isasymlongrightarrow \ k\ \isasymle \ y)\isanewline  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1334  | 
\ 2.\ \isasymAnd x.\ \isasymlbrakk P\ k;\ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\ \isasymle \ x;\ P\ x\ \isasymand \ (\isasymforall y.\ P\ y\ \isasymlongrightarrow \ x\ \isasymle \ y)\isasymrbrakk \isanewline  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1335  | 
\ \ \ \ \ \ \ \ \isasymLongrightarrow \ x\ =\ k%  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1336  | 
\end{isabelle}
 | 
| 11458 | 1337  | 
As always with \isa{the_equality}, we must show existence and
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1338  | 
uniqueness of the claimed solution,~\isa{k}.  Existence, the first
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1339  | 
subgoal, is trivial. Uniqueness, the second subgoal, follows by antisymmetry:  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1340  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1341  | 
\isasymlbrakk x\ \isasymle \ y;\ y\ \isasymle \ x\isasymrbrakk \ \isasymLongrightarrow \ x\ =\ y%  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1342  | 
\rulename{order_antisym}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1343  | 
\end{isabelle}
 | 
| 11458 | 1344  | 
The assumptions imply both \isa{k~\isasymle~x} and \isa{x~\isasymle~k}.  One
 | 
1345  | 
call to \isa{auto} does it all: 
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1346  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1347  | 
\isacommand{by}\ (auto\ intro:\ order_antisym)
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1348  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1349  | 
|
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1350  | 
|
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1351  | 
\subsection{Indefinite Descriptions}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1352  | 
|
| 11458 | 1353  | 
\index{Hilbert's $\varepsilon$-operator}%
 | 
| 11077 | 1354  | 
\index{descriptions!indefinite}%
 | 
| 11458 | 1355  | 
An indefinite description is traditionally written $\varepsilon x. P(x)$ and is  | 
1356  | 
known as Hilbert's $\varepsilon$-operator. It denotes  | 
|
1357  | 
some $x$ such that $P(x)$ is true, provided one exists.  | 
|
1358  | 
Isabelle uses \sdx{SOME} for the Greek letter~$\varepsilon$.
 | 
|
1359  | 
||
| 33057 | 1360  | 
Here is the definition of~\cdx{inv},\footnote{In fact, \isa{inv} is defined via a second constant \isa{inv_into}, which we ignore here.} which expresses inverses of
 | 
| 11417 | 1361  | 
functions:  | 
| 11077 | 1362  | 
\begin{isabelle}
 | 
1363  | 
inv\ f\ \isasymequiv \ \isasymlambda y.\ SOME\ x.\ f\ x\ =\ y%  | 
|
1364  | 
\rulename{inv_def}
 | 
|
1365  | 
\end{isabelle}
 | 
|
| 11458 | 1366  | 
Using \isa{SOME} rather than \isa{THE} makes \isa{inv~f} behave well
 | 
1367  | 
even if \isa{f} is not injective.  As it happens, most useful theorems about
 | 
|
1368  | 
\isa{inv} do assume the function to be injective.
 | 
|
1369  | 
||
| 11406 | 1370  | 
The inverse of \isa{f}, when applied to \isa{y}, returns some~\isa{x} such that
 | 
| 11077 | 1371  | 
\isa{f~x~=~y}.  For example, we can prove \isa{inv~Suc} really is the inverse
 | 
1372  | 
of the \isa{Suc} function 
 | 
|
1373  | 
\begin{isabelle}
 | 
|
1374  | 
\isacommand{lemma}\ "inv\ Suc\ (Suc\ n)\ =\ n"\isanewline
 | 
|
1375  | 
\isacommand{by}\ (simp\ add:\ inv_def)
 | 
|
1376  | 
\end{isabelle}
 | 
|
1377  | 
||
1378  | 
\noindent  | 
|
1379  | 
The proof is a one-liner: the subgoal simplifies to a degenerate application of  | 
|
1380  | 
\isa{SOME}, which is then erased.  In detail, the left-hand side simplifies
 | 
|
1381  | 
to \isa{SOME\ x.\ Suc\ x\ =\ Suc\ n}, then to \isa{SOME\ x.\ x\ =\ n} and
 | 
|
1382  | 
finally to~\isa{n}.  
 | 
|
1383  | 
||
1384  | 
We know nothing about what  | 
|
1385  | 
\isa{inv~Suc} returns when applied to zero.  The proof above still treats
 | 
|
1386  | 
\isa{SOME} as a definite description, since it only reasons about
 | 
|
| 11458 | 1387  | 
situations in which the value is described uniquely.  Indeed, \isa{SOME}
 | 
1388  | 
satisfies this rule:  | 
|
1389  | 
\begin{isabelle}
 | 
|
1390  | 
\isasymlbrakk P\ a;\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow \ x\ =\ a\isasymrbrakk \  | 
|
1391  | 
\isasymLongrightarrow \ (SOME\ x.\ P\ x)\ =\ a%  | 
|
1392  | 
\rulenamedx{some_equality}
 | 
|
1393  | 
\end{isabelle}
 | 
|
1394  | 
To go further is  | 
|
| 11077 | 1395  | 
tricky and requires rules such as these:  | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1396  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1397  | 
P\ x\ \isasymLongrightarrow \ P\ (SOME\ x.\ P\ x)  | 
| 11417 | 1398  | 
\rulenamedx{someI}\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1399  | 
\isasymlbrakk P\ a;\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow \ Q\  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1400  | 
x\isasymrbrakk \ \isasymLongrightarrow \ Q\ (SOME\ x.\ P\ x)  | 
| 11417 | 1401  | 
\rulenamedx{someI2}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1402  | 
\end{isabelle}
 | 
| 11406 | 1403  | 
Rule \isa{someI} is basic: if anything satisfies \isa{P} then so does
 | 
1404  | 
\hbox{\isa{SOME\ x.\ P\ x}}.  The repetition of~\isa{P} in the conclusion makes it
 | 
|
1405  | 
difficult to apply in a backward proof, so the derived rule \isa{someI2} is
 | 
|
1406  | 
also provided.  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1407  | 
|
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1408  | 
\medskip  | 
| 11406 | 1409  | 
For example, let us prove the \rmindex{axiom of choice}:
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1410  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1411  | 
\isacommand{theorem}\ axiom_of_choice:
 | 
| 11077 | 1412  | 
\ "(\isasymforall x.\ \isasymexists y.\ P\ x\ y)\ \isasymLongrightarrow \  | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1413  | 
\isasymexists f.\ \isasymforall x.\ P\ x\ (f\ x)"\isanewline  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1414  | 
\isacommand{apply}\ (rule\ exI,\ rule\ allI)\isanewline
 | 
| 10971 | 1415  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1416  | 
\ 1.\ \isasymAnd x.\ \isasymforall x.\ \isasymexists y.\ P\ x\ y\  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1417  | 
\isasymLongrightarrow \ P\ x\ (?f\ x)  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1418  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1419  | 
%  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1420  | 
We have applied the introduction rules; now it is time to apply the elimination  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1421  | 
rules.  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1422  | 
|
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1423  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1424  | 
\isacommand{apply}\ (drule\ spec,\ erule\ exE)\isanewline
 | 
| 10971 | 1425  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1426  | 
\ 1.\ \isasymAnd x\ y.\ P\ (?x2\ x)\ y\ \isasymLongrightarrow \ P\ x\ (?f\ x)  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1427  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1428  | 
|
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1429  | 
\noindent  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1430  | 
The rule \isa{someI} automatically instantiates
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1431  | 
\isa{f} to \hbox{\isa{\isasymlambda x.\ SOME y.\ P\ x\ y}}, which is the choice
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1432  | 
function.  It also instantiates \isa{?x2\ x} to \isa{x}.
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1433  | 
\begin{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1434  | 
\isacommand{by}\ (rule\ someI)\isanewline
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1435  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1436  | 
|
| 11077 | 1437  | 
\subsubsection{Historical Note}
 | 
1438  | 
The original purpose of Hilbert's $\varepsilon$-operator was to express an  | 
|
1439  | 
existential destruction rule:  | 
|
1440  | 
\[ \infer{P[(\varepsilon x. P) / \, x]}{\exists x.\,P} \]
 | 
|
1441  | 
This rule is seldom used for that purpose --- it can cause exponential  | 
|
1442  | 
blow-up --- but it is occasionally used as an introduction rule  | 
|
| 13791 | 1443  | 
for the~$\varepsilon$-operator.  Its name in HOL is \tdxbold{someI_ex}.%%
 | 
| 11458 | 1444  | 
\index{description operators|)}
 | 
| 11077 | 1445  | 
|
1446  | 
||
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1447  | 
\section{Some Proofs That Fail}
 | 
| 10295 | 1448  | 
|
| 11077 | 1449  | 
\index{proofs!examples of failing|(}%
 | 
| 10295 | 1450  | 
Most of the examples in this tutorial involve proving theorems. But not every  | 
1451  | 
conjecture is true, and it can be instructive to see how  | 
|
1452  | 
proofs fail. Here we attempt to prove a distributive law involving  | 
|
1453  | 
the existential quantifier and conjunction.  | 
|
1454  | 
\begin{isabelle}
 | 
|
| 11077 | 1455  | 
\isacommand{lemma}\ "({\isasymexists}x.\ P\ x)\ \isasymand\ 
 | 
1456  | 
({\isasymexists}x.\ Q\ x)\ \isasymLongrightarrow\ {\isasymexists}x.\ P\ x\
 | 
|
1457  | 
\isasymand\ Q\ x"  | 
|
1458  | 
\end{isabelle}
 | 
|
1459  | 
The first steps are routine. We apply conjunction elimination to break  | 
|
1460  | 
the assumption into two existentially quantified assumptions.  | 
|
1461  | 
Applying existential elimination removes one of the quantifiers.  | 
|
1462  | 
\begin{isabelle}
 | 
|
| 10295 | 1463  | 
\isacommand{apply}\ (erule\ conjE)\isanewline
 | 
1464  | 
\isacommand{apply}\ (erule\ exE)\isanewline
 | 
|
| 10596 | 1465  | 
\ 1.\ \isasymAnd x.\ \isasymlbrakk{\isasymexists}x.\ Q\ x;\ P\ x\isasymrbrakk\ \isasymLongrightarrow\ {\isasymexists}x.\ P\ x\ \isasymand\ Q\ x
 | 
| 10295 | 1466  | 
\end{isabelle}
 | 
1467  | 
%  | 
|
1468  | 
When we remove the other quantifier, we get a different bound  | 
|
1469  | 
variable in the subgoal.  (The name \isa{xa} is generated automatically.)
 | 
|
1470  | 
\begin{isabelle}
 | 
|
| 11077 | 1471  | 
\isacommand{apply}\ (erule\ exE)\isanewline
 | 
| 10596 | 1472  | 
\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\  | 
| 10295 | 1473  | 
\isasymLongrightarrow\ {\isasymexists}x.\ P\ x\ \isasymand\ Q\ x
 | 
1474  | 
\end{isabelle}
 | 
|
1475  | 
The proviso of the existential elimination rule has forced the variables to  | 
|
1476  | 
differ: we can hardly expect two arbitrary values to be equal! There is  | 
|
1477  | 
no way to prove this subgoal. Removing the  | 
|
1478  | 
conclusion's existential quantifier yields two  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1479  | 
identical placeholders, which can become  any term involving the variables \isa{x}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1480  | 
and~\isa{xa}.  We need one to become \isa{x}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1481  | 
and the other to become~\isa{xa}, but Isabelle requires all instances of a
 | 
| 10295 | 1482  | 
placeholder to be identical.  | 
1483  | 
\begin{isabelle}
 | 
|
| 11077 | 1484  | 
\isacommand{apply}\ (rule\ exI)\isanewline
 | 
1485  | 
\isacommand{apply}\ (rule\ conjI)\isanewline
 | 
|
| 10596 | 1486  | 
\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\  | 
| 10295 | 1487  | 
\isasymLongrightarrow\ P\ (?x3\ x\ xa)\isanewline  | 
| 10596 | 1488  | 
\ 2.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\ \isasymLongrightarrow\ Q\ (?x3\ x\ xa)  | 
| 10295 | 1489  | 
\end{isabelle}
 | 
1490  | 
We can prove either subgoal  | 
|
1491  | 
using the \isa{assumption} method.  If we prove the first one, the placeholder
 | 
|
| 11077 | 1492  | 
changes into~\isa{x}. 
 | 
| 10295 | 1493  | 
\begin{isabelle}
 | 
| 11077 | 1494  | 
\ \isacommand{apply}\ assumption\isanewline
 | 
| 10596 | 1495  | 
\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\  | 
| 10295 | 1496  | 
\isasymLongrightarrow\ Q\ x  | 
1497  | 
\end{isabelle}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1498  | 
We are left with a subgoal that cannot be proved.  Applying the \isa{assumption}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1499  | 
method results in an error message:  | 
| 10295 | 1500  | 
\begin{isabelle}
 | 
1501  | 
*** empty result sequence -- proof command failed  | 
|
1502  | 
\end{isabelle}
 | 
|
| 11077 | 1503  | 
When interacting with Isabelle via the shell interface,  | 
1504  | 
you can abandon a proof using the \isacommand{oops} command.
 | 
|
| 10295 | 1505  | 
|
1506  | 
\medskip  | 
|
1507  | 
||
1508  | 
Here is another abortive proof, illustrating the interaction between  | 
|
1509  | 
bound variables and unknowns.  | 
|
1510  | 
If $R$ is a reflexive relation,  | 
|
1511  | 
is there an $x$ such that $R\,x\,y$ holds for all $y$? Let us see what happens when  | 
|
1512  | 
we attempt to prove it.  | 
|
1513  | 
\begin{isabelle}
 | 
|
| 11406 | 1514  | 
\isacommand{lemma}\ "\isasymforall y.\ R\ y\ y\ \isasymLongrightarrow 
 | 
1515  | 
\ \isasymexists x.\ \isasymforall y.\ R\ x\ y"  | 
|
| 11077 | 1516  | 
\end{isabelle}
 | 
1517  | 
First, we remove the existential quantifier. The new proof state has an  | 
|
1518  | 
unknown, namely~\isa{?x}. 
 | 
|
1519  | 
\begin{isabelle}
 | 
|
| 10295 | 1520  | 
\isacommand{apply}\ (rule\ exI)\isanewline
 | 
| 11077 | 1521  | 
\ 1.\ \isasymforall y.\ R\ y\ y\ \isasymLongrightarrow \ \isasymforall y.\ R\ ?x\ y%  | 
| 10295 | 1522  | 
\end{isabelle}
 | 
| 11077 | 1523  | 
It looks like we can just apply \isa{assumption}, but it fails.  Isabelle
 | 
1524  | 
refuses to substitute \isa{y}, a bound variable, for~\isa{?x}; that would be
 | 
|
1525  | 
a bound variable capture. We can still try to finish the proof in some  | 
|
1526  | 
other way. We remove the universal quantifier from the conclusion, moving  | 
|
1527  | 
the bound variable~\isa{y} into the subgoal.  But note that it is still
 | 
|
1528  | 
bound!  | 
|
| 10295 | 1529  | 
\begin{isabelle}
 | 
| 11077 | 1530  | 
\isacommand{apply}\ (rule\ allI)\isanewline
 | 
1531  | 
\ 1.\ \isasymAnd y.\ \isasymforall y.\ R\ y\ y\ \isasymLongrightarrow \ R\ ?x\ y%  | 
|
| 10295 | 1532  | 
\end{isabelle}
 | 
1533  | 
Finally, we try to apply our reflexivity assumption. We obtain a  | 
|
1534  | 
new assumption whose identical placeholders may be replaced by  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1535  | 
any term involving~\isa{y}. 
 | 
| 10295 | 1536  | 
\begin{isabelle}
 | 
| 11077 | 1537  | 
\isacommand{apply}\ (drule\ spec)\isanewline
 | 
| 10596 | 1538  | 
\ 1.\ \isasymAnd y.\ R\ (?z2\ y)\ (?z2\ y)\ \isasymLongrightarrow\ R\ ?x\ y  | 
| 10295 | 1539  | 
\end{isabelle}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1540  | 
This subgoal can only be proved by putting \isa{y} for all the placeholders,
 | 
| 11077 | 1541  | 
making the assumption and conclusion become \isa{R\ y\ y}.  Isabelle can
 | 
1542  | 
replace \isa{?z2~y} by \isa{y}; this involves instantiating
 | 
|
1543  | 
\isa{?z2} to the identity function.  But, just as two steps earlier,
 | 
|
1544  | 
Isabelle refuses to substitute~\isa{y} for~\isa{?x}.
 | 
|
| 10295 | 1545  | 
This example is typical of how Isabelle enforces sound quantifier reasoning.  | 
| 11077 | 1546  | 
\index{proofs!examples of failing|)}
 | 
| 10295 | 1547  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1548  | 
\section{Proving Theorems Using the {\tt\slshape blast} Method}
 | 
| 10295 | 1549  | 
|
| 11077 | 1550  | 
\index{*blast (method)|(}%
 | 
| 11406 | 1551  | 
It is hard to prove many theorems using the methods  | 
1552  | 
described above. A proof may be hundreds of steps long. You  | 
|
| 10295 | 1553  | 
may need to search among different ways of proving certain  | 
1554  | 
subgoals. Often a choice that proves one subgoal renders another  | 
|
1555  | 
impossible to prove. There are further complications that we have not  | 
|
1556  | 
discussed, concerning negation and disjunction. Isabelle's  | 
|
| 11406 | 1557  | 
\textbf{classical reasoner} is a family of tools that perform such
 | 
| 10295 | 1558  | 
proofs automatically. The most important of these is the  | 
| 10596 | 1559  | 
\isa{blast} method. 
 | 
| 10295 | 1560  | 
|
1561  | 
In this section, we shall first see how to use the classical  | 
|
1562  | 
reasoner in its default mode and then how to insert additional  | 
|
1563  | 
rules, enabling it to work in new problem domains.  | 
|
1564  | 
||
1565  | 
We begin with examples from pure predicate logic. The following  | 
|
1566  | 
example is known as Andrew's challenge. Peter Andrews designed  | 
|
| 11406 | 1567  | 
it to be hard to prove by automatic means.  | 
1568  | 
It is particularly hard for a resolution prover, where  | 
|
1569  | 
converting the nested biconditionals to  | 
|
1570  | 
clause form produces a combinatorial  | 
|
1571  | 
explosion~\cite{pelletier86}. However, the
 | 
|
| 11077 | 1572  | 
\isa{blast} method proves it in a fraction  of a second. 
 | 
| 10295 | 1573  | 
\begin{isabelle}
 | 
1574  | 
\isacommand{lemma}\
 | 
|
1575  | 
"(({\isasymexists}x.\
 | 
|
1576  | 
{\isasymforall}y.\
 | 
|
| 10301 | 1577  | 
p(x){=}p(y))\
 | 
| 10295 | 1578  | 
=\  | 
1579  | 
(({\isasymexists}x.\
 | 
|
| 10301 | 1580  | 
q(x))=({\isasymforall}y.\
 | 
1581  | 
p(y))))\  | 
|
| 10295 | 1582  | 
\ \ =\ \ \ \ \isanewline  | 
1583  | 
\ \ \ \ \ \ \ \  | 
|
1584  | 
(({\isasymexists}x.\
 | 
|
1585  | 
{\isasymforall}y.\
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1586  | 
q(x){=}q(y))\ =\ (({\isasymexists}x.\ p(x))=({\isasymforall}y.\ q(y))))"\isanewline
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1587  | 
\isacommand{by}\ blast
 | 
| 10295 | 1588  | 
\end{isabelle}
 | 
1589  | 
The next example is a logic problem composed by Lewis Carroll.  | 
|
| 10596 | 1590  | 
The \isa{blast} method finds it trivial. Moreover, it turns out 
 | 
| 11406 | 1591  | 
that not all of the assumptions are necessary. We can  | 
| 10295 | 1592  | 
experiment with variations of this formula and see which ones  | 
1593  | 
can be proved.  | 
|
1594  | 
\begin{isabelle}
 | 
|
1595  | 
\isacommand{lemma}\
 | 
|
1596  | 
"({\isasymforall}x.\
 | 
|
1597  | 
honest(x)\ \isasymand\  | 
|
1598  | 
industrious(x)\ \isasymlongrightarrow\  | 
|
| 10301 | 1599  | 
healthy(x))\  | 
| 10295 | 1600  | 
\isasymand\ \ \isanewline  | 
1601  | 
\ \ \ \ \ \ \ \ \isasymnot\ ({\isasymexists}x.\
 | 
|
1602  | 
grocer(x)\ \isasymand\  | 
|
| 10301 | 1603  | 
healthy(x))\  | 
| 10295 | 1604  | 
\isasymand\ \isanewline  | 
1605  | 
\ \ \ \ \ \ \ \ ({\isasymforall}x.\
 | 
|
1606  | 
industrious(x)\ \isasymand\  | 
|
1607  | 
grocer(x)\ \isasymlongrightarrow\  | 
|
| 10301 | 1608  | 
honest(x))\  | 
| 10295 | 1609  | 
\isasymand\ \isanewline  | 
1610  | 
\ \ \ \ \ \ \ \ ({\isasymforall}x.\
 | 
|
1611  | 
cyclist(x)\ \isasymlongrightarrow\  | 
|
| 10301 | 1612  | 
industrious(x))\  | 
| 10295 | 1613  | 
\isasymand\ \isanewline  | 
1614  | 
\ \ \ \ \ \ \ \ ({\isasymforall}x.\
 | 
|
1615  | 
{\isasymnot}healthy(x)\ \isasymand\
 | 
|
1616  | 
cyclist(x)\ \isasymlongrightarrow\  | 
|
| 10301 | 1617  | 
{\isasymnot}honest(x))\
 | 
| 10295 | 1618  | 
\ \isanewline  | 
1619  | 
\ \ \ \ \ \ \ \ \isasymlongrightarrow\  | 
|
1620  | 
({\isasymforall}x.\
 | 
|
1621  | 
grocer(x)\ \isasymlongrightarrow\  | 
|
| 10301 | 1622  | 
{\isasymnot}cyclist(x))"\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1623  | 
\isacommand{by}\ blast
 | 
| 10295 | 1624  | 
\end{isabelle}
 | 
| 10596 | 1625  | 
The \isa{blast} method is also effective for set theory, which is
 | 
| 11406 | 1626  | 
described in the next chapter. The formula below may look horrible, but  | 
1627  | 
the \isa{blast} method proves it in milliseconds. 
 | 
|
| 10295 | 1628  | 
\begin{isabelle}
 | 
| 10301 | 1629  | 
\isacommand{lemma}\ "({\isasymUnion}i{\isasymin}I.\ A(i))\ \isasyminter\ ({\isasymUnion}j{\isasymin}J.\ B(j))\ =\isanewline
 | 
1630  | 
\ \ \ \ \ \ \ \ ({\isasymUnion}i{\isasymin}I.\ {\isasymUnion}j{\isasymin}J.\ A(i)\ \isasyminter\ B(j))"\isanewline
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1631  | 
\isacommand{by}\ blast
 | 
| 10295 | 1632  | 
\end{isabelle}
 | 
1633  | 
||
1634  | 
Few subgoals are couched purely in predicate logic and set theory.  | 
|
1635  | 
We can extend the scope of the classical reasoner by giving it new rules.  | 
|
1636  | 
Extending it effectively requires understanding the notions of  | 
|
1637  | 
introduction, elimination and destruction rules. Moreover, there is a  | 
|
| 11077 | 1638  | 
distinction between safe and unsafe rules. A  | 
1639  | 
\textbf{safe}\indexbold{safe rules} rule is one that can be applied 
 | 
|
1640  | 
backwards without losing information; an  | 
|
1641  | 
\textbf{unsafe}\indexbold{unsafe rules} rule loses  information, perhaps
 | 
|
1642  | 
transforming the subgoal into one that cannot be proved. The safe/unsafe  | 
|
| 10295 | 1643  | 
distinction affects the proof search: if a proof attempt fails, the  | 
1644  | 
classical reasoner backtracks to the most recent unsafe rule application  | 
|
1645  | 
and makes another choice.  | 
|
1646  | 
||
1647  | 
An important special case avoids all these complications. A logical  | 
|
1648  | 
equivalence, which in higher-order logic is an equality between  | 
|
1649  | 
formulas, can be given to the classical  | 
|
| 11406 | 1650  | 
reasoner and simplifier by using the attribute \attrdx{iff}.  You 
 | 
| 10295 | 1651  | 
should do so if the right hand side of the equivalence is  | 
1652  | 
simpler than the left-hand side.  | 
|
1653  | 
||
1654  | 
For example, here is a simple fact about list concatenation.  | 
|
1655  | 
The result of appending two lists is empty if and only if both  | 
|
1656  | 
of the lists are themselves empty. Obviously, applying this equivalence  | 
|
1657  | 
will result in a simpler goal. When stating this lemma, we include  | 
|
| 11406 | 1658  | 
the \attrdx{iff} attribute. Once we have proved the lemma, Isabelle 
 | 
| 10295 | 1659  | 
will make it known to the classical reasoner (and to the simplifier).  | 
1660  | 
\begin{isabelle}
 | 
|
1661  | 
\isacommand{lemma}\
 | 
|
| 10854 | 1662  | 
[iff]:\ "(xs{\isacharat}ys\ =\ [])\ =\
 | 
| 10971 | 1663  | 
(xs=[]\ \isasymand\ ys=[])"\isanewline  | 
| 10854 | 1664  | 
\isacommand{apply}\ (induct_tac\ xs)\isanewline
 | 
1665  | 
\isacommand{apply}\ (simp_all)\isanewline
 | 
|
| 10295 | 1666  | 
\isacommand{done}
 | 
1667  | 
\end{isabelle}
 | 
|
1668  | 
%  | 
|
1669  | 
This fact about multiplication is also appropriate for  | 
|
| 11406 | 1670  | 
the \attrdx{iff} attribute:
 | 
| 10295 | 1671  | 
\begin{isabelle}
 | 
| 10596 | 1672  | 
(\mbox{?m}\ *\ \mbox{?n}\ =\ 0)\ =\ (\mbox{?m}\ =\ 0\ \isasymor\ \mbox{?n}\ =\ 0)
 | 
| 10295 | 1673  | 
\end{isabelle}
 | 
1674  | 
A product is zero if and only if one of the factors is zero. The  | 
|
| 10971 | 1675  | 
reasoning involves a disjunction. Proving new rules for  | 
| 10295 | 1676  | 
disjunctive reasoning is hard, but translating to an actual disjunction  | 
1677  | 
works: the classical reasoner handles disjunction properly.  | 
|
1678  | 
||
| 11406 | 1679  | 
In more detail, this is how the \attrdx{iff} attribute works.  It converts
 | 
| 10295 | 1680  | 
the equivalence $P=Q$ to a pair of rules: the introduction  | 
1681  | 
rule $Q\Imp P$ and the destruction rule $P\Imp Q$. It gives both to the  | 
|
1682  | 
classical reasoner as safe rules, ensuring that all occurrences of $P$ in  | 
|
1683  | 
a subgoal are replaced by~$Q$. The simplifier performs the same  | 
|
1684  | 
replacement, since \isa{iff} gives $P=Q$ to the
 | 
|
| 11406 | 1685  | 
simplifier.  | 
1686  | 
||
1687  | 
Classical reasoning is different from  | 
|
1688  | 
simplification. Simplification is deterministic. It applies rewrite rules  | 
|
1689  | 
repeatedly, as long as possible, transforming a goal into another goal. Classical  | 
|
1690  | 
reasoning uses search and backtracking in order to prove a goal outright.%  | 
|
| 11077 | 1691  | 
\index{*blast (method)|)}%
 | 
1692  | 
||
| 10295 | 1693  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1694  | 
\section{Other Classical Reasoning Methods}
 | 
| 10295 | 1695  | 
|
| 10596 | 1696  | 
The \isa{blast} method is our main workhorse for proving theorems 
 | 
| 10295 | 1697  | 
automatically. Other components of the classical reasoner interact  | 
1698  | 
with the simplifier. Still others perform classical reasoning  | 
|
1699  | 
to a limited extent, giving the user fine control over the proof.  | 
|
1700  | 
||
| 11077 | 1701  | 
Of the latter methods, the most useful is  | 
| 11406 | 1702  | 
\methdx{clarify}.
 | 
| 11077 | 1703  | 
It performs  | 
| 10295 | 1704  | 
all obvious reasoning steps without splitting the goal into multiple  | 
| 10971 | 1705  | 
parts. It does not apply unsafe rules that could render the  | 
1706  | 
goal unprovable. By performing the obvious  | 
|
| 11077 | 1707  | 
steps, \isa{clarify} lays bare the difficult parts of the problem, 
 | 
| 10295 | 1708  | 
where human intervention is necessary.  | 
1709  | 
||
1710  | 
For example, the following conjecture is false:  | 
|
1711  | 
\begin{isabelle}
 | 
|
1712  | 
\isacommand{lemma}\ "({\isasymforall}x.\ P\ x)\ \isasymand\
 | 
|
1713  | 
({\isasymexists}x.\ Q\ x)\ \isasymlongrightarrow\ ({\isasymforall}x.\ P\ x\
 | 
|
1714  | 
\isasymand\ Q\ x)"\isanewline  | 
|
1715  | 
\isacommand{apply}\ clarify
 | 
|
1716  | 
\end{isabelle}
 | 
|
| 11077 | 1717  | 
The \isa{blast} method would simply fail, but \isa{clarify} presents 
 | 
| 10295 | 1718  | 
a subgoal that helps us see why we cannot continue the proof.  | 
1719  | 
\begin{isabelle}
 | 
|
| 10596 | 1720  | 
\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk{\isasymforall}x.\ P\ x;\ Q\
 | 
| 10295 | 1721  | 
xa\isasymrbrakk\ \isasymLongrightarrow\ P\ x\ \isasymand\ Q\ x  | 
1722  | 
\end{isabelle}
 | 
|
1723  | 
The proof must fail because the assumption \isa{Q\ xa} and conclusion \isa{Q\ x}
 | 
|
1724  | 
refer to distinct bound variables.  To reach this state, \isa{clarify} applied
 | 
|
1725  | 
the introduction rules for \isa{\isasymlongrightarrow} and \isa{\isasymforall}
 | 
|
| 12535 | 1726  | 
and the elimination rule for \isa{\isasymand}.  It did not apply the introduction
 | 
| 10295 | 1727  | 
rule for  \isa{\isasymand} because of its policy never to split goals.
 | 
1728  | 
||
| 11406 | 1729  | 
Also available is \methdx{clarsimp}, a method
 | 
1730  | 
that interleaves \isa{clarify} and \isa{simp}.  Also there is  \methdx{safe},
 | 
|
1731  | 
which like \isa{clarify} performs obvious steps but even applies those that
 | 
|
| 11077 | 1732  | 
split goals.  | 
| 10295 | 1733  | 
|
| 11406 | 1734  | 
The \methdx{force} method applies the classical
 | 
| 11077 | 1735  | 
reasoner and simplifier to one goal.  | 
| 10295 | 1736  | 
Unless it can prove the goal, it fails. Contrast  | 
| 10546 | 1737  | 
that with the \isa{auto} method, which also combines classical reasoning 
 | 
| 10295 | 1738  | 
with simplification. The latter's purpose is to prove all the  | 
1739  | 
easy subgoals and parts of subgoals. Unfortunately, it can produce  | 
|
1740  | 
large numbers of new subgoals; also, since it proves some subgoals  | 
|
1741  | 
and splits others, it obscures the structure of the proof tree.  | 
|
| 10546 | 1742  | 
The \isa{force} method does not have these drawbacks. Another 
 | 
1743  | 
difference: \isa{force} tries harder than {\isa{auto}} to prove 
 | 
|
| 10295 | 1744  | 
its goal, so it can take much longer to terminate.  | 
1745  | 
||
1746  | 
Older components of the classical reasoner have largely been  | 
|
| 10596 | 1747  | 
superseded by \isa{blast}, but they still have niche applications. 
 | 
1748  | 
Most important among these are \isa{fast} and \isa{best}. While \isa{blast} 
 | 
|
| 10295 | 1749  | 
searches for proofs using a built-in first-order reasoner, these  | 
1750  | 
earlier methods search for proofs using standard Isabelle inference.  | 
|
| 
11179
 
bee6673b020a
subst method and a new section on rule, rule_tac, etc
 
paulson 
parents: 
11159 
diff
changeset
 | 
1751  | 
That makes them slower but enables them to work in the  | 
| 10295 | 1752  | 
presence of the more unusual features of Isabelle rules, such  | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1753  | 
as type classes and function unknowns. For example, recall the introduction rule  | 
| 10971 | 1754  | 
for Hilbert's $\varepsilon$-operator:  | 
| 10295 | 1755  | 
\begin{isabelle}
 | 
| 10546 | 1756  | 
?P\ ?x\ \isasymLongrightarrow\ ?P\ (SOME\ x.\ ?P x)  | 
| 10295 | 1757  | 
\rulename{someI}
 | 
1758  | 
\end{isabelle}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1759  | 
%  | 
| 10295 | 1760  | 
The repeated occurrence of the variable \isa{?P} makes this rule tricky 
 | 
1761  | 
to apply. Consider this contrived example:  | 
|
1762  | 
\begin{isabelle}
 | 
|
| 10596 | 1763  | 
\isacommand{lemma}\ "\isasymlbrakk Q\ a;\ P\ a\isasymrbrakk\isanewline
 | 
| 10295 | 1764  | 
\ \ \ \ \ \ \ \ \,\isasymLongrightarrow\ P\ (SOME\ x.\ P\ x\ \isasymand\ Q\ x)\  | 
1765  | 
\isasymand\ Q\ (SOME\ x.\ P\ x\ \isasymand\ Q\ x)"\isanewline  | 
|
1766  | 
\isacommand{apply}\ (rule\ someI)
 | 
|
1767  | 
\end{isabelle}
 | 
|
1768  | 
%  | 
|
1769  | 
We can apply rule \isa{someI} explicitly.  It yields the 
 | 
|
1770  | 
following subgoal:  | 
|
1771  | 
\begin{isabelle}
 | 
|
| 10596 | 1772  | 
\ 1.\ \isasymlbrakk Q\ a;\ P\ a\isasymrbrakk\ \isasymLongrightarrow\ P\ ?x\  | 
| 10295 | 1773  | 
\isasymand\ Q\ ?x%  | 
1774  | 
\end{isabelle}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1775  | 
The proof from this point is trivial. Could we have  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1776  | 
proved the theorem with a single command? Not using \isa{blast}: it
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1777  | 
cannot perform the higher-order unification needed here. The  | 
| 11406 | 1778  | 
\methdx{fast} method succeeds: 
 | 
| 10295 | 1779  | 
\begin{isabelle}
 | 
1780  | 
\isacommand{apply}\ (fast\ intro!:\ someI)
 | 
|
1781  | 
\end{isabelle}
 | 
|
1782  | 
||
| 11406 | 1783  | 
The \methdx{best} method is similar to
 | 
| 11077 | 1784  | 
\isa{fast} but it uses a  best-first search instead of depth-first search.
 | 
1785  | 
Accordingly, it is slower but is less susceptible to divergence.  | 
|
| 11406 | 1786  | 
Transitivity  rules usually cause \isa{fast} to loop where \isa{best} 
 | 
1787  | 
can often manage.  | 
|
| 10295 | 1788  | 
|
1789  | 
Here is a summary of the classical reasoning methods:  | 
|
1790  | 
\begin{itemize}
 | 
|
| 11406 | 1791  | 
\item \methdx{blast} works automatically and is the fastest
 | 
1792  | 
||
1793  | 
\item \methdx{clarify} and \methdx{clarsimp} perform obvious steps without
 | 
|
1794  | 
splitting the goal;  \methdx{safe} even splits goals
 | 
|
1795  | 
||
1796  | 
\item \methdx{force} uses classical reasoning and simplification to prove a goal;
 | 
|
1797  | 
 \methdx{auto} is similar but leaves what it cannot prove
 | 
|
1798  | 
||
1799  | 
\item \methdx{fast} and \methdx{best} are legacy methods that work well with rules
 | 
|
1800  | 
involving unusual features  | 
|
| 10295 | 1801  | 
\end{itemize}
 | 
1802  | 
A table illustrates the relationships among four of these methods.  | 
|
1803  | 
\begin{center}
 | 
|
1804  | 
\begin{tabular}{r|l|l|}
 | 
|
1805  | 
& no split & split \\ \hline  | 
|
| 11406 | 1806  | 
  no simp  & \methdx{clarify}    & \methdx{safe} \\ \hline
 | 
1807  | 
     simp  & \methdx{clarsimp}   & \methdx{auto} \\ \hline
 | 
|
| 10295 | 1808  | 
\end{tabular}
 | 
1809  | 
\end{center}
 | 
|
1810  | 
||
| 16546 | 1811  | 
\section{Finding More Theorems}
 | 
1812  | 
\label{sec:find2}
 | 
|
1813  | 
\input{Rules/document/find2.tex}
 | 
|
1814  | 
||
| 10295 | 1815  | 
|
| 11406 | 1816  | 
\section{Forward Proof: Transforming Theorems}\label{sec:forward}
 | 
| 10295 | 1817  | 
|
| 11077 | 1818  | 
\index{forward proof|(}%
 | 
| 10295 | 1819  | 
Forward proof means deriving new facts from old ones. It is the  | 
1820  | 
most fundamental type of proof. Backward proof, by working from goals to  | 
|
1821  | 
subgoals, can help us find a difficult proof. But it is  | 
|
| 14403 | 1822  | 
not always the best way of presenting the proof thus found. Forward  | 
| 10301 | 1823  | 
proof is particularly good for reasoning from the general  | 
| 11406 | 1824  | 
to the specific. For example, consider this distributive law for  | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1825  | 
the greatest common divisor:  | 
| 10295 | 1826  | 
\[ k\times\gcd(m,n) = \gcd(k\times m,k\times n)\]  | 
1827  | 
||
1828  | 
Putting $m=1$ we get (since $\gcd(1,n)=1$ and $k\times1=k$)  | 
|
1829  | 
\[ k = \gcd(k,k\times n)\]  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1830  | 
We have derived a new fact; if re-oriented, it might be  | 
| 10295 | 1831  | 
useful for simplification. After re-orienting it and putting $n=1$, we  | 
1832  | 
derive another useful law:  | 
|
1833  | 
\[ \gcd(k,k)=k \]  | 
|
1834  | 
Substituting values for variables --- instantiation --- is a forward step.  | 
|
1835  | 
Re-orientation works by applying the symmetry of equality to  | 
|
1836  | 
an equation, so it too is a forward step.  | 
|
1837  | 
||
| 14403 | 1838  | 
\subsection{Modifying a Theorem using {\tt\slshape of},  {\tt\slshape where}
 | 
1839  | 
 and {\tt\slshape THEN}}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1840  | 
|
| 15952 | 1841  | 
\label{sec:THEN}
 | 
1842  | 
||
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1843  | 
Let us reproduce our examples in Isabelle. Recall that in  | 
| 25258 | 1844  | 
{\S}\ref{sec:fun-simplification} we declared the recursive function
 | 
| 11406 | 1845  | 
\isa{gcd}:\index{*gcd (constant)|(}
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1846  | 
\begin{isabelle}
 | 
| 25264 | 1847  | 
\isacommand{fun}\ gcd\ ::\ "nat\ \isasymRightarrow \ nat\ \isasymRightarrow \ nat"\ \isakeyword{where}\isanewline
 | 
1848  | 
\ \ "gcd\ m\ n\ =\ (if\ n=0\ then\ m\ else\ gcd\ n\ (m\ mod\ n))"  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1849  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1850  | 
%  | 
| 12333 | 1851  | 
From this definition, it is possible to prove the distributive law.  | 
1852  | 
That takes us to the starting point for our example.  | 
|
| 11077 | 1853  | 
\begin{isabelle}
 | 
| 25264 | 1854  | 
?k\ *\ gcd\ ?m\ ?n\ =\ gcd\ (?k\ *\ ?m)\ (?k\ *\ ?n)  | 
| 10295 | 1855  | 
\rulename{gcd_mult_distrib2}
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1856  | 
\end{isabelle}
 | 
| 11406 | 1857  | 
%  | 
1858  | 
The first step in our derivation is to replace \isa{?m} by~1.  We instantiate the
 | 
|
1859  | 
theorem using~\attrdx{of}, which identifies variables in order of their
 | 
|
1860  | 
appearance from left to right.  In this case, the variables  are \isa{?k}, \isa{?m}
 | 
|
1861  | 
and~\isa{?n}. So, the expression
 | 
|
| 10295 | 1862  | 
\hbox{\texttt{[of k 1]}} replaces \isa{?k} by~\isa{k} and \isa{?m}
 | 
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
1863  | 
by~\isa{1}.
 | 
| 10295 | 1864  | 
\begin{isabelle}
 | 
1865  | 
\isacommand{lemmas}\ gcd_mult_0\ =\ gcd_mult_distrib2\ [of\ k\ 1]
 | 
|
1866  | 
\end{isabelle}
 | 
|
1867  | 
%  | 
|
| 11406 | 1868  | 
The keyword \commdx{lemmas} declares a new theorem, which can be derived
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1869  | 
from an existing one using attributes such as \isa{[of~k~1]}.
 | 
| 10295 | 1870  | 
The command  | 
1871  | 
\isa{thm gcd_mult_0}
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1872  | 
displays the result:  | 
| 10295 | 1873  | 
\begin{isabelle}
 | 
| 25264 | 1874  | 
\ \ \ \ \ k\ *\ gcd\ 1\ ?n\ =\ gcd\ (k\ *\ 1)\ (k\ *\ ?n)  | 
| 10295 | 1875  | 
\end{isabelle}
 | 
| 14403 | 1876  | 
Something is odd: \isa{k} is an ordinary variable, while \isa{?n} 
 | 
| 10295 | 1877  | 
is schematic. We did not specify an instantiation  | 
| 14403 | 1878  | 
for \isa{?n}.  In its present form, the theorem does not allow 
 | 
1879  | 
substitution for \isa{k}.  One solution is to avoid giving an instantiation for
 | 
|
| 10295 | 1880  | 
\isa{?k}: instead of a term we can put an underscore~(\isa{_}).  For example,
 | 
1881  | 
\begin{isabelle}
 | 
|
1882  | 
\ \ \ \ \ gcd_mult_distrib2\ [of\ _\ 1]  | 
|
1883  | 
\end{isabelle}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1884  | 
replaces \isa{?m} by~\isa{1} but leaves \isa{?k} unchanged.  
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1885  | 
|
| 14403 | 1886  | 
An equivalent solution is to use the attribute \isa{where}. 
 | 
1887  | 
\begin{isabelle}
 | 
|
1888  | 
\ \ \ \ \ gcd\_mult\_distrib2\ [where\ m=1]  | 
|
1889  | 
\end{isabelle}
 | 
|
1890  | 
While \isa{of} refers to
 | 
|
1891  | 
variables by their position, \isa{where} refers to variables by name. Multiple
 | 
|
1892  | 
instantiations are separated by~\isa{and}, as in this example:
 | 
|
1893  | 
\begin{isabelle}
 | 
|
1894  | 
\ \ \ \ \ gcd\_mult\_distrib2\ [where\ m=1\ and\ k=1]  | 
|
1895  | 
\end{isabelle}
 | 
|
1896  | 
||
1897  | 
We now continue the present example with the version of \isa{gcd_mult_0}
 | 
|
1898  | 
shown above, which has \isa{k} instead of \isa{?k}.
 | 
|
1899  | 
Once we have replaced \isa{?m} by~1, we must next simplify
 | 
|
1900  | 
the theorem \isa{gcd_mult_0}, performing the steps 
 | 
|
| 11406 | 1901  | 
$\gcd(1,n)=1$ and $k\times1=k$.  The \attrdx{simplified}
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1902  | 
attribute takes a theorem  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1903  | 
and returns the result of simplifying it, with respect to the default  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1904  | 
simplification rules:  | 
| 10295 | 1905  | 
\begin{isabelle}
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1906  | 
\isacommand{lemmas}\ gcd_mult_1\ =\ gcd_mult_0\
 | 
| 10295 | 1907  | 
[simplified]%  | 
1908  | 
\end{isabelle}
 | 
|
1909  | 
%  | 
|
1910  | 
Again, we display the resulting theorem:  | 
|
1911  | 
\begin{isabelle}
 | 
|
| 25264 | 1912  | 
\ \ \ \ \ k\ =\ gcd\ k\ (k\ *\ ?n)  | 
| 10295 | 1913  | 
\end{isabelle}
 | 
1914  | 
%  | 
|
1915  | 
To re-orient the equation requires the symmetry rule:  | 
|
1916  | 
\begin{isabelle}
 | 
|
1917  | 
?s\ =\ ?t\  | 
|
1918  | 
\isasymLongrightarrow\ ?t\ =\  | 
|
1919  | 
?s%  | 
|
| 11417 | 1920  | 
\rulenamedx{sym}
 | 
| 10295 | 1921  | 
\end{isabelle}
 | 
1922  | 
The following declaration gives our equation to \isa{sym}:
 | 
|
1923  | 
\begin{isabelle}
 | 
|
| 11077 | 1924  | 
\ \ \ \isacommand{lemmas}\ gcd_mult\ =\ gcd_mult_1\ [THEN\ sym]
 | 
| 10295 | 1925  | 
\end{isabelle}
 | 
1926  | 
%  | 
|
1927  | 
Here is the result:  | 
|
1928  | 
\begin{isabelle}
 | 
|
| 25264 | 1929  | 
\ \ \ \ \ gcd\ k\ (k\ *\ ?n)\ =\ k%  | 
| 10295 | 1930  | 
\end{isabelle}
 | 
| 11406 | 1931  | 
\isa{THEN~sym}\indexbold{*THEN (attribute)} gives the current theorem to the
 | 
| 11077 | 1932  | 
rule \isa{sym} and returns the resulting conclusion.  The effect is to
 | 
1933  | 
exchange the two operands of the equality. Typically \isa{THEN} is used
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1934  | 
with destruction rules.  Also useful is \isa{THEN~spec}, which removes the
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1935  | 
quantifier from a theorem of the form $\forall x.\,P$, and \isa{THEN~mp},
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1936  | 
which converts the implication $P\imp Q$ into the rule  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1937  | 
$\vcenter{\infer{Q}{P}}$. Similar to \isa{mp} are the following two rules,
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1938  | 
which extract the two directions of reasoning about a boolean equivalence:  | 
| 10295 | 1939  | 
\begin{isabelle}
 | 
| 10596 | 1940  | 
\isasymlbrakk?Q\ =\ ?P;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%  | 
| 11417 | 1941  | 
\rulenamedx{iffD1}%
 | 
| 10295 | 1942  | 
\isanewline  | 
| 10596 | 1943  | 
\isasymlbrakk?P\ =\ ?Q;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%  | 
| 11417 | 1944  | 
\rulenamedx{iffD2}
 | 
| 10295 | 1945  | 
\end{isabelle}
 | 
1946  | 
%  | 
|
1947  | 
Normally we would never name the intermediate theorems  | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
1948  | 
such as \isa{gcd_mult_0} and \isa{gcd_mult_1} but would combine
 | 
| 10295 | 1949  | 
the three forward steps:  | 
1950  | 
\begin{isabelle}
 | 
|
1951  | 
\isacommand{lemmas}\ gcd_mult\ =\ gcd_mult_distrib2\ [of\ k\ 1,\ simplified,\ THEN\ sym]%
 | 
|
1952  | 
\end{isabelle}
 | 
|
1953  | 
The directives, or attributes, are processed from left to right. This  | 
|
1954  | 
declaration of \isa{gcd_mult} is equivalent to the
 | 
|
1955  | 
previous one.  | 
|
1956  | 
||
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1957  | 
Such declarations can make the proof script hard to read. Better  | 
| 10295 | 1958  | 
is to state the new lemma explicitly and to prove it using a single  | 
1959  | 
\isa{rule} method whose operand is expressed using forward reasoning:
 | 
|
1960  | 
\begin{isabelle}
 | 
|
| 25264 | 1961  | 
\isacommand{lemma}\ gcd\_mult\ [simp]:\ "gcd\ k\ (k*n)\ =\ k"\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1962  | 
\isacommand{by}\ (rule\ gcd_mult_distrib2\ [of\ k\ 1,\ simplified,\ THEN\ sym])
 | 
| 10295 | 1963  | 
\end{isabelle}
 | 
1964  | 
Compared with the previous proof of \isa{gcd_mult}, this
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1965  | 
version shows the reader what has been proved. Also, the result will be processed  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1966  | 
in the normal way. In particular, Isabelle generalizes over all variables: the  | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1967  | 
resulting theorem will have {\isa{?k}} instead of {\isa{k}}.
 | 
| 10295 | 1968  | 
|
1969  | 
At the start of this section, we also saw a proof of $\gcd(k,k)=k$. Here  | 
|
| 11406 | 1970  | 
is the Isabelle version:\index{*gcd (constant)|)}
 | 
| 10295 | 1971  | 
\begin{isabelle}
 | 
| 25264 | 1972  | 
\isacommand{lemma}\ gcd\_self\ [simp]:\ "gcd\ k\ k\ =\ k"\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1973  | 
\isacommand{by}\ (rule\ gcd_mult\ [of\ k\ 1,\ simplified])
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1974  | 
\end{isabelle}
 | 
| 
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1975  | 
|
| 11406 | 1976  | 
\begin{warn}
 | 
| 12535 | 1977  | 
To give~\isa{of} a nonatomic term, enclose it in quotation marks, as in
 | 
| 11406 | 1978  | 
\isa{[of "k*m"]}.  The term must not contain unknowns: an
 | 
1979  | 
attribute such as  | 
|
1980  | 
\isa{[of "?k*m"]} will be rejected.
 | 
|
1981  | 
\end{warn}
 | 
|
1982  | 
||
| 15952 | 1983  | 
%Answer is now included in that section! Is a modified version of this  | 
1984  | 
% exercise worth including? E.g. find a difference between the two ways  | 
|
1985  | 
% of substituting.  | 
|
1986  | 
%\begin{exercise}
 | 
|
1987  | 
%In {\S}\ref{sec:subst} the method \isa{subst\ mult_commute} was applied.  How
 | 
|
1988  | 
%can we achieve the same effect using \isa{THEN} with the rule \isa{ssubst}?
 | 
|
1989  | 
%% answer rule (mult_commute [THEN ssubst])  | 
|
1990  | 
%\end{exercise}
 | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
1991  | 
|
| 11406 | 1992  | 
\subsection{Modifying a Theorem using {\tt\slshape OF}}
 | 
| 10295 | 1993  | 
|
| 11406 | 1994  | 
\index{*OF (attribute)|(}%
 | 
| 11077 | 1995  | 
Recall that \isa{of} generates an instance of a
 | 
1996  | 
rule by specifying values for its variables.  Analogous is \isa{OF}, which
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1997  | 
generates an instance of a rule by specifying facts for its premises.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
1998  | 
|
| 11406 | 1999  | 
We again need the divides relation\index{divides relation} of number theory, which
 | 
2000  | 
as we recall is defined by  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2001  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2002  | 
?m\ dvd\ ?n\ \isasymequiv\ {\isasymexists}k.\ ?n\ =\ ?m\ *\ k
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2003  | 
\rulename{dvd_def}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2004  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2005  | 
%  | 
| 12333 | 2006  | 
Suppose, for example, that we have proved the following rule.  | 
2007  | 
It states that if $k$ and $n$ are relatively prime  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2008  | 
and if $k$ divides $m\times n$ then $k$ divides $m$.  | 
| 10295 | 2009  | 
\begin{isabelle}
 | 
| 25264 | 2010  | 
\isasymlbrakk gcd ?k ?n {=} 1;\ ?k\ dvd\ ?m * ?n\isasymrbrakk\
 | 
| 10295 | 2011  | 
\isasymLongrightarrow\ ?k\ dvd\ ?m  | 
2012  | 
\rulename{relprime_dvd_mult}
 | 
|
2013  | 
\end{isabelle}
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2014  | 
We can use \isa{OF} to create an instance of this rule.
 | 
| 10295 | 2015  | 
First, we  | 
2016  | 
prove an instance of its first premise:  | 
|
2017  | 
\begin{isabelle}
 | 
|
| 25264 | 2018  | 
\isacommand{lemma}\ relprime\_20\_81:\ "gcd\ 20\ 81\ =\ 1"\isanewline
 | 
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
2019  | 
\isacommand{by}\ (simp\ add:\ gcd.simps)
 | 
| 10295 | 2020  | 
\end{isabelle}
 | 
2021  | 
We have evaluated an application of the \isa{gcd} function by
 | 
|
| 11077 | 2022  | 
simplification. Expression evaluation involving recursive functions is not  | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2023  | 
guaranteed to terminate, and it can be slow; Isabelle  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2024  | 
performs arithmetic by rewriting symbolic bit strings. Here,  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2025  | 
however, the simplification takes less than one second. We can  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2026  | 
give this new lemma to \isa{OF}.  The expression
 | 
| 10295 | 2027  | 
\begin{isabelle}
 | 
2028  | 
\ \ \ \ \ relprime_dvd_mult [OF relprime_20_81]  | 
|
2029  | 
\end{isabelle}
 | 
|
2030  | 
yields the theorem  | 
|
2031  | 
\begin{isabelle}
 | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2032  | 
\ \ \ \ \ 20\ dvd\ (?m\ *\ 81)\ \isasymLongrightarrow\ 20\ dvd\ ?m%  | 
| 10295 | 2033  | 
\end{isabelle}
 | 
2034  | 
%  | 
|
| 10596 | 2035  | 
\isa{OF} takes any number of operands.  Consider 
 | 
| 10295 | 2036  | 
the following facts about the divides relation:  | 
2037  | 
\begin{isabelle}
 | 
|
2038  | 
\isasymlbrakk?k\ dvd\ ?m;\  | 
|
2039  | 
?k\ dvd\ ?n\isasymrbrakk\  | 
|
2040  | 
\isasymLongrightarrow\ ?k\ dvd\  | 
|
| 10971 | 2041  | 
?m\ +\ ?n  | 
| 10295 | 2042  | 
\rulename{dvd_add}\isanewline
 | 
2043  | 
?m\ dvd\ ?m%  | 
|
2044  | 
\rulename{dvd_refl}
 | 
|
2045  | 
\end{isabelle}
 | 
|
2046  | 
Let us supply \isa{dvd_refl} for each of the premises of \isa{dvd_add}:
 | 
|
2047  | 
\begin{isabelle}
 | 
|
2048  | 
\ \ \ \ \ dvd_add [OF dvd_refl dvd_refl]  | 
|
2049  | 
\end{isabelle}
 | 
|
2050  | 
Here is the theorem that we have expressed:  | 
|
2051  | 
\begin{isabelle}
 | 
|
| 10596 | 2052  | 
\ \ \ \ \ ?k\ dvd\ (?k\ +\ ?k)  | 
| 10295 | 2053  | 
\end{isabelle}
 | 
2054  | 
As with \isa{of}, we can use the \isa{_} symbol to leave some positions
 | 
|
2055  | 
unspecified:  | 
|
2056  | 
\begin{isabelle}
 | 
|
2057  | 
\ \ \ \ \ dvd_add [OF _ dvd_refl]  | 
|
2058  | 
\end{isabelle}
 | 
|
2059  | 
The result is  | 
|
2060  | 
\begin{isabelle}
 | 
|
| 10971 | 2061  | 
\ \ \ \ \ ?k\ dvd\ ?m\ \isasymLongrightarrow\ ?k\ dvd\ ?m\ +\ ?k  | 
| 10295 | 2062  | 
\end{isabelle}
 | 
2063  | 
||
| 10596 | 2064  | 
You may have noticed that \isa{THEN} and \isa{OF} are based on 
 | 
| 10295 | 2065  | 
the same idea, namely to combine two rules. They differ in the  | 
2066  | 
order of the combination and thus in their effect.  We use \isa{THEN}
 | 
|
2067  | 
typically with a destruction rule to extract a subformula of the current  | 
|
2068  | 
theorem.  We use \isa{OF} with a list of facts to generate an instance of
 | 
|
| 11077 | 2069  | 
the current theorem.%  | 
| 11406 | 2070  | 
\index{*OF (attribute)|)}
 | 
| 10295 | 2071  | 
|
| 
10848
 
7b3ee4695fe6
various changes including the SOME examples, rule_format and "by"
 
paulson 
parents: 
10792 
diff
changeset
 | 
2072  | 
Here is a summary of some primitives for forward reasoning:  | 
| 10295 | 2073  | 
\begin{itemize}
 | 
| 11406 | 2074  | 
\item \attrdx{of} instantiates the variables of a rule to a list of terms
 | 
2075  | 
\item \attrdx{OF} applies a rule to a list of theorems
 | 
|
2076  | 
\item \attrdx{THEN} gives a theorem to a named rule and returns the
 | 
|
| 10295 | 2077  | 
conclusion  | 
| 11406 | 2078  | 
%\item \attrdx{rule_format} puts a theorem into standard form
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2079  | 
%  by removing \isa{\isasymlongrightarrow} and~\isa{\isasymforall}
 | 
| 11406 | 2080  | 
\item \attrdx{simplified} applies the simplifier to a theorem
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2081  | 
\item \isacommand{lemmas} assigns a name to the theorem produced by the
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2082  | 
attributes above  | 
| 10295 | 2083  | 
\end{itemize}
 | 
2084  | 
||
2085  | 
||
| 11406 | 2086  | 
\section{Forward Reasoning in a Backward Proof}
 | 
| 10295 | 2087  | 
|
| 10967 | 2088  | 
We have seen that the forward proof directives work well within a backward  | 
| 11077 | 2089  | 
proof. There are many ways to achieve a forward style using our existing  | 
2090  | 
proof methods. We shall also meet some new methods that perform forward  | 
|
2091  | 
reasoning.  | 
|
| 10967 | 2092  | 
|
2093  | 
The methods \isa{drule}, \isa{frule}, \isa{drule_tac}, etc.,
 | 
|
2094  | 
reason forward from a subgoal. We have seen them already, using rules such as  | 
|
2095  | 
\isa{mp} and
 | 
|
2096  | 
\isa{spec} to operate on formulae.  They can also operate on terms, using rules
 | 
|
2097  | 
such as these:  | 
|
2098  | 
\begin{isabelle}
 | 
|
2099  | 
x\ =\ y\ \isasymLongrightarrow \ f\ x\ =\ f\ y%  | 
|
| 11417 | 2100  | 
\rulenamedx{arg_cong}\isanewline
 | 
| 10967 | 2101  | 
i\ \isasymle \ j\ \isasymLongrightarrow \ i\ *\ k\ \isasymle \ j\ *\ k%  | 
2102  | 
\rulename{mult_le_mono1}
 | 
|
2103  | 
\end{isabelle}
 | 
|
2104  | 
||
2105  | 
For example, let us prove a fact about divisibility in the natural numbers:  | 
|
2106  | 
\begin{isabelle}
 | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2107  | 
\isacommand{lemma}\ "2\ \isasymle \ u\ \isasymLongrightarrow \ u*m\ \isasymnoteq
 | 
| 10967 | 2108  | 
\ Suc(u*n)"\isanewline  | 
| 12408 | 2109  | 
\isacommand{apply}\ (intro\ notI)\isanewline
 | 
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2110  | 
\ 1.\ \isasymlbrakk 2\ \isasymle \ u;\ u\ *\ m\ =\ Suc\ (u\ *\ n)\isasymrbrakk \ \isasymLongrightarrow \ False%  | 
| 10967 | 2111  | 
\end{isabelle}
 | 
2112  | 
%  | 
|
2113  | 
The key step is to apply the function \ldots\isa{mod\ u} to both sides of the
 | 
|
2114  | 
equation  | 
|
| 11077 | 2115  | 
\isa{u*m\ =\ Suc(u*n)}:\index{*drule_tac (method)}
 | 
| 10967 | 2116  | 
\begin{isabelle}
 | 
2117  | 
\isacommand{apply}\ (drule_tac\ f="\isasymlambda x.\ x\ mod\ u"\ \isakeyword{in}\
 | 
|
2118  | 
arg_cong)\isanewline  | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2119  | 
\ 1.\ \isasymlbrakk 2\ \isasymle \ u;\ u\ *\ m\ mod\ u\ =\ Suc\ (u\ *\ n)\ mod\  | 
| 10967 | 2120  | 
u\isasymrbrakk \ \isasymLongrightarrow \ False  | 
2121  | 
\end{isabelle}
 | 
|
2122  | 
%  | 
|
2123  | 
Simplification reduces the left side to 0 and the right side to~1, yielding the  | 
|
2124  | 
required contradiction.  | 
|
2125  | 
\begin{isabelle}
 | 
|
2126  | 
\isacommand{apply}\ (simp\ add:\ mod_Suc)\isanewline
 | 
|
2127  | 
\isacommand{done}
 | 
|
2128  | 
\end{isabelle}
 | 
|
2129  | 
||
2130  | 
Our proof has used a fact about remainder:  | 
|
2131  | 
\begin{isabelle}
 | 
|
2132  | 
Suc\ m\ mod\ n\ =\isanewline  | 
|
2133  | 
(if\ Suc\ (m\ mod\ n)\ =\ n\ then\ 0\ else\ Suc\ (m\ mod\ n))  | 
|
2134  | 
\rulename{mod_Suc}
 | 
|
2135  | 
\end{isabelle}
 | 
|
2136  | 
||
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2137  | 
\subsection{The Method {\tt\slshape insert}}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2138  | 
|
| 11406 | 2139  | 
\index{*insert (method)|(}%
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2140  | 
The \isa{insert} method
 | 
| 
30649
 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 
nipkow 
parents: 
27167 
diff
changeset
 | 
2141  | 
inserts a given theorem as a new assumption of all subgoals. This  | 
| 11077 | 2142  | 
already is a forward step; moreover, we may (as always when using a  | 
2143  | 
theorem) apply  | 
|
| 10596 | 2144  | 
\isa{of}, \isa{THEN} and other directives.  The new assumption can then
 | 
| 
30649
 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 
nipkow 
parents: 
27167 
diff
changeset
 | 
2145  | 
be used to help prove the subgoals.  | 
| 10295 | 2146  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2147  | 
For example, consider this theorem about the divides relation. The first  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2148  | 
proof step inserts the distributive law for  | 
| 10295 | 2149  | 
\isa{gcd}.  We specify its variables as shown. 
 | 
2150  | 
\begin{isabelle}
 | 
|
| 25264 | 2151  | 
\isacommand{lemma}\ relprime\_dvd\_mult:\ \isanewline
 | 
2152  | 
\ \ \ \ \ \ "\isasymlbrakk \ gcd\ k\ n\ =\ 1;\ k\ dvd\ m*n\ \isasymrbrakk \ \isasymLongrightarrow \ k\ dvd\ m"\isanewline  | 
|
2153  | 
\isacommand{apply}\ (insert\ gcd_mult_distrib2\ [of\ m\ k\ n])
 | 
|
| 10295 | 2154  | 
\end{isabelle}
 | 
2155  | 
In the resulting subgoal, note how the equation has been  | 
|
2156  | 
inserted:  | 
|
2157  | 
\begin{isabelle}
 | 
|
| 25264 | 2158  | 
\ 1.\ \isasymlbrakk gcd\ k\ n\ =\ 1;\ k\ dvd\ m\ *\ n;\ m\ *\ gcd\ k\ n\ =\ gcd\ (m\ *\ k)\ (m\ *\ n)\isasymrbrakk \isanewline  | 
2159  | 
\isaindent{\ 1.\ }\isasymLongrightarrow \ k\ dvd\ m%
 | 
|
| 10295 | 2160  | 
\end{isabelle}
 | 
| 25264 | 2161  | 
The next proof step utilizes the assumption \isa{gcd\ k\ n\ =\ 1}
 | 
2162  | 
(note that \isa{Suc\ 0} is another expression for 1):
 | 
|
| 10295 | 2163  | 
\begin{isabelle}
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2164  | 
\isacommand{apply}(simp)\isanewline
 | 
| 25264 | 2165  | 
\ 1.\ \isasymlbrakk gcd\ k\ n\ =\ Suc\ 0;\ k\ dvd\ m\ *\ n;\ m\ =\ gcd\ (m\ *\ k)\ (m\ *\ n)\isasymrbrakk \isanewline  | 
2166  | 
\isaindent{\ 1.\ }\isasymLongrightarrow \ k\ dvd\ m%
 | 
|
| 10295 | 2167  | 
\end{isabelle}
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2168  | 
Simplification has yielded an equation for~\isa{m}.  The rest of the proof
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2169  | 
is omitted.  | 
| 10295 | 2170  | 
|
2171  | 
\medskip  | 
|
| 11417 | 2172  | 
Here is another demonstration of \isa{insert}.  Division and
 | 
2173  | 
remainder obey a well-known law:  | 
|
| 10295 | 2174  | 
\begin{isabelle}
 | 
| 10596 | 2175  | 
(?m\ div\ ?n)\ *\ ?n\ +\ ?m\ mod\ ?n\ =\ ?m  | 
| 10295 | 2176  | 
\rulename{mod_div_equality}
 | 
2177  | 
\end{isabelle}
 | 
|
2178  | 
||
2179  | 
We refer to this law explicitly in the following proof:  | 
|
2180  | 
\begin{isabelle}
 | 
|
2181  | 
\isacommand{lemma}\ div_mult_self_is_m:\ \isanewline
 | 
|
| 10596 | 2182  | 
\ \ \ \ \ \ "0{\isacharless}n\ \isasymLongrightarrow\ (m*n)\ div\ n\ =\
 | 
2183  | 
(m::nat)"\isanewline  | 
|
2184  | 
\isacommand{apply}\ (insert\ mod_div_equality\ [of\ "m*n"\ n])\isanewline
 | 
|
| 10295 | 2185  | 
\isacommand{apply}\ (simp)\isanewline
 | 
2186  | 
\isacommand{done}
 | 
|
2187  | 
\end{isabelle}
 | 
|
2188  | 
The first step inserts the law, specifying \isa{m*n} and
 | 
|
| 10301 | 2189  | 
\isa{n} for its variables.  Notice that non-trivial expressions must be
 | 
| 10295 | 2190  | 
enclosed in quotation marks. Here is the resulting  | 
2191  | 
subgoal, with its new assumption:  | 
|
2192  | 
\begin{isabelle}
 | 
|
2193  | 
%0\ \isacharless\ n\ \isasymLongrightarrow\ (m\  | 
|
| 10596 | 2194  | 
%*\ n)\ div\ n\ =\ m\isanewline  | 
| 10295 | 2195  | 
\ 1.\ \isasymlbrakk0\ \isacharless\  | 
| 10596 | 2196  | 
n;\ \ (m\ *\ n)\ div\ n\ *\ n\ +\ (m\ *\ n)\ mod\ n\  | 
2197  | 
=\ m\ *\ n\isasymrbrakk\isanewline  | 
|
2198  | 
\ \ \ \ \isasymLongrightarrow\ (m\ *\ n)\ div\ n\  | 
|
| 10295 | 2199  | 
=\ m  | 
2200  | 
\end{isabelle}
 | 
|
| 10596 | 2201  | 
Simplification reduces \isa{(m\ *\ n)\ mod\ n} to zero.
 | 
| 10295 | 2202  | 
Then it cancels the factor~\isa{n} on both
 | 
| 11406 | 2203  | 
sides of the equation \isa{(m\ *\ n)\ div\ n\ *\ n\ =\ m\ *\ n}, proving the
 | 
2204  | 
theorem.  | 
|
2205  | 
||
2206  | 
\begin{warn}
 | 
|
2207  | 
Any unknowns in the theorem given to \methdx{insert} will be universally
 | 
|
2208  | 
quantified in the new assumption.  | 
|
2209  | 
\end{warn}%
 | 
|
2210  | 
\index{*insert (method)|)}
 | 
|
| 10295 | 2211  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2212  | 
\subsection{The Method {\tt\slshape subgoal_tac}}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2213  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2214  | 
\index{*subgoal_tac (method)}%
 | 
| 11406 | 2215  | 
A related method is \isa{subgoal_tac}, but instead
 | 
| 11077 | 2216  | 
of inserting a theorem as an assumption, it inserts an arbitrary formula.  | 
| 10295 | 2217  | 
This formula must be proved later as a separate subgoal. The  | 
2218  | 
idea is to claim that the formula holds on the basis of the current  | 
|
2219  | 
assumptions, to use this claim to complete the proof, and finally  | 
|
| 11406 | 2220  | 
to justify the claim. It gives the proof  | 
2221  | 
some structure. If you find yourself generating a complex assumption by a  | 
|
2222  | 
long series of forward steps, consider using \isa{subgoal_tac} instead: you can
 | 
|
2223  | 
state the formula you are aiming for, and perhaps prove it automatically.  | 
|
| 10295 | 2224  | 
|
2225  | 
Look at the following example.  | 
|
2226  | 
\begin{isabelle}
 | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2227  | 
\isacommand{lemma}\ "\isasymlbrakk(z::int)\ <\ 37;\ 66\ <\ 2*z;\ z*z\
 | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2228  | 
\isasymnoteq\ 1225;\ Q(34);\ Q(36)\isasymrbrakk\isanewline  | 
| 10295 | 2229  | 
\ \ \ \ \ \ \ \ \,\isasymLongrightarrow\ Q(z)"\isanewline  | 
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2230  | 
\isacommand{apply}\ (subgoal_tac\ "z\ =\ 34\ \isasymor\ z\ =\
 | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2231  | 
36")\isanewline  | 
| 10295 | 2232  | 
\isacommand{apply}\ blast\isanewline
 | 
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2233  | 
\isacommand{apply}\ (subgoal_tac\ "z\ \isasymnoteq\ 35")\isanewline
 | 
| 10295 | 2234  | 
\isacommand{apply}\ arith\isanewline
 | 
2235  | 
\isacommand{apply}\ force\isanewline
 | 
|
2236  | 
\isacommand{done}
 | 
|
2237  | 
\end{isabelle}
 | 
|
| 11406 | 2238  | 
The first assumption tells us  | 
2239  | 
that \isa{z} is no greater than~36. The second tells us that \isa{z} 
 | 
|
2240  | 
is at least~34. The third assumption tells us that \isa{z} cannot be 35, since
 | 
|
2241  | 
$35\times35=1225$.  So \isa{z} is either 34 or~36, and since \isa{Q} holds for
 | 
|
| 10295 | 2242  | 
both of those values, we have the conclusion.  | 
2243  | 
||
2244  | 
The Isabelle proof closely follows this reasoning. The first  | 
|
2245  | 
step is to claim that \isa{z} is either 34 or 36. The resulting proof 
 | 
|
2246  | 
state gives us two subgoals:  | 
|
2247  | 
\begin{isabelle}
 | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2248  | 
%\isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2249  | 
%Q\ 34;\ Q\ 36\isasymrbrakk\ \isasymLongrightarrow\ Q\ z\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2250  | 
\ 1.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\ Q\ 34;\ Q\ 36;\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2251  | 
\ \ \ \ \ z\ =\ 34\ \isasymor\ z\ =\ 36\isasymrbrakk\isanewline  | 
| 10295 | 2252  | 
\ \ \ \ \isasymLongrightarrow\ Q\ z\isanewline  | 
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2253  | 
\ 2.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\ Q\ 34;\ Q\ 36\isasymrbrakk\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2254  | 
\ \ \ \ \isasymLongrightarrow\ z\ =\ 34\ \isasymor\ z\ =\ 36  | 
| 10295 | 2255  | 
\end{isabelle}
 | 
| 10971 | 2256  | 
The first subgoal is trivial (\isa{blast}), but for the second Isabelle needs help to eliminate
 | 
| 10295 | 2257  | 
the case  | 
| 10596 | 2258  | 
\isa{z}=35.  The second invocation  of {\isa{subgoal_tac}} leaves two
 | 
| 10295 | 2259  | 
subgoals:  | 
2260  | 
\begin{isabelle}
 | 
|
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2261  | 
\ 1.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2262  | 
1225;\ Q\ 34;\ Q\ 36;\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2263  | 
\ \ \ \ \ z\ \isasymnoteq\ 35\isasymrbrakk\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2264  | 
\ \ \ \ \isasymLongrightarrow\ z\ =\ 34\ \isasymor\ z\ =\ 36\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2265  | 
\ 2.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\ Q\ 34;\ Q\ 36\isasymrbrakk\isanewline  | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2266  | 
\ \ \ \ \isasymLongrightarrow\ z\ \isasymnoteq\ 35  | 
| 10295 | 2267  | 
\end{isabelle}
 | 
2268  | 
||
| 10971 | 2269  | 
Assuming that \isa{z} is not 35, the first subgoal follows by linear arithmetic
 | 
2270  | 
(\isa{arith}). For the second subgoal we apply the method \isa{force}, 
 | 
|
| 10295 | 2271  | 
which proceeds by assuming that \isa{z}=35 and arriving at a contradiction.
 | 
2272  | 
||
2273  | 
||
2274  | 
\medskip  | 
|
2275  | 
Summary of these methods:  | 
|
2276  | 
\begin{itemize}
 | 
|
| 11406 | 2277  | 
\item \methdx{insert} adds a theorem as a new assumption
 | 
2278  | 
\item \methdx{subgoal_tac} adds a formula as a new assumption and leaves the
 | 
|
| 10295 | 2279  | 
subgoal of proving that formula  | 
2280  | 
\end{itemize}
 | 
|
| 11077 | 2281  | 
\index{forward proof|)}
 | 
| 10967 | 2282  | 
|
2283  | 
||
2284  | 
\section{Managing Large Proofs}
 | 
|
2285  | 
||
2286  | 
Naturally you should try to divide proofs into manageable parts. Look for lemmas  | 
|
2287  | 
that can be proved separately. Sometimes you will observe that they are  | 
|
2288  | 
instances of much simpler facts. On other occasions, no lemmas suggest themselves  | 
|
2289  | 
and you are forced to cope with a long proof involving many subgoals.  | 
|
2290  | 
||
2291  | 
\subsection{Tacticals, or Control Structures}
 | 
|
2292  | 
||
| 11406 | 2293  | 
\index{tacticals|(}%
 | 
| 10967 | 2294  | 
If the proof is long, perhaps it at least has some regularity. Then you can  | 
| 11406 | 2295  | 
express it more concisely using \textbf{tacticals}, which provide control
 | 
| 10967 | 2296  | 
structures. Here is a proof (it would be a one-liner using  | 
2297  | 
\isa{blast}, but forget that) that contains a series of repeated
 | 
|
2298  | 
commands:  | 
|
2299  | 
%  | 
|
2300  | 
\begin{isabelle}
 | 
|
2301  | 
\isacommand{lemma}\ "\isasymlbrakk P\isasymlongrightarrow Q;\
 | 
|
2302  | 
Q\isasymlongrightarrow R;\ R\isasymlongrightarrow S;\ P\isasymrbrakk \  | 
|
2303  | 
\isasymLongrightarrow \ S"\isanewline  | 
|
2304  | 
\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
 | 
|
2305  | 
\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
 | 
|
2306  | 
\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
 | 
|
2307  | 
\isacommand{apply}\ (assumption)\isanewline
 | 
|
2308  | 
\isacommand{done}
 | 
|
2309  | 
\end{isabelle}
 | 
|
2310  | 
%  | 
|
2311  | 
Each of the three identical commands finds an implication and proves its  | 
|
2312  | 
antecedent by assumption.  The first one finds \isa{P\isasymlongrightarrow Q} and
 | 
|
2313  | 
\isa{P}, concluding~\isa{Q}; the second one concludes~\isa{R} and the third one
 | 
|
2314  | 
concludes~\isa{S}.  The final step matches the assumption \isa{S} with the goal to
 | 
|
2315  | 
be proved.  | 
|
2316  | 
||
| 11406 | 2317  | 
Suffixing a method with a plus sign~(\isa+)\index{*"+ (tactical)}
 | 
| 10967 | 2318  | 
expresses one or more repetitions:  | 
2319  | 
\begin{isabelle}
 | 
|
2320  | 
\isacommand{lemma}\ "\isasymlbrakk P\isasymlongrightarrow Q;\ Q\isasymlongrightarrow R;\ R\isasymlongrightarrow S;\ P\isasymrbrakk \ \isasymLongrightarrow \ S"\isanewline
 | 
|
2321  | 
\isacommand{by}\ (drule\ mp,\ assumption)+
 | 
|
2322  | 
\end{isabelle}
 | 
|
2323  | 
%  | 
|
2324  | 
Using \isacommand{by} takes care of the final use of \isa{assumption}.  The new
 | 
|
2325  | 
proof is more concise. It is also more general: the repetitive method works  | 
|
2326  | 
for a chain of implications having any length, not just three.  | 
|
2327  | 
||
2328  | 
Choice is another control structure. Separating two methods by a vertical  | 
|
| 11406 | 2329  | 
% we must use ?? rather than "| as the sorting item because somehow the presence  | 
2330  | 
% of | (even quoted) stops hyperref from putting |hyperpage at the end of the index  | 
|
2331  | 
% entry.  | 
|
2332  | 
bar~(\isa|)\index{??@\texttt{"|} (tactical)}  gives the effect of applying the
 | 
|
2333  | 
first method, and if that fails, trying the second. It can be combined with  | 
|
2334  | 
repetition, when the choice must be made over and over again. Here is a chain of  | 
|
2335  | 
implications in which most of the antecedents are proved by assumption, but one is  | 
|
2336  | 
proved by arithmetic:  | 
|
| 10967 | 2337  | 
\begin{isabelle}
 | 
| 
12156
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2338  | 
\isacommand{lemma}\ "\isasymlbrakk Q\isasymlongrightarrow R;\ P\isasymlongrightarrow Q;\ x<5\isasymlongrightarrow P;\
 | 
| 
 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 
paulson 
parents: 
11494 
diff
changeset
 | 
2339  | 
Suc\ x\ <\ 5\isasymrbrakk \ \isasymLongrightarrow \ R"\ \isanewline  | 
| 10967 | 2340  | 
\isacommand{by}\ (drule\ mp,\ (assumption|arith))+
 | 
2341  | 
\end{isabelle}
 | 
|
2342  | 
The \isa{arith}
 | 
|
2343  | 
method can prove $x<5$ from $x+1<5$, but it cannot duplicate the effect of  | 
|
2344  | 
\isa{assumption}.  Therefore, we combine these methods using the choice
 | 
|
2345  | 
operator.  | 
|
2346  | 
||
| 11406 | 2347  | 
A postfixed question mark~(\isa?)\index{*"? (tactical)} expresses zero or one
 | 
2348  | 
repetitions of a method. It can also be viewed as the choice between executing a  | 
|
| 12540 | 2349  | 
method and doing nothing. It is useless at top level but can be valuable  | 
2350  | 
within other control structures; for example,  | 
|
2351  | 
\isa{($m$+)?} performs zero or more repetitions of method~$m$.%
 | 
|
| 11406 | 2352  | 
\index{tacticals|)}
 | 
| 10967 | 2353  | 
|
2354  | 
||
2355  | 
\subsection{Subgoal Numbering}
 | 
|
2356  | 
||
2357  | 
Another problem in large proofs is contending with huge  | 
|
2358  | 
subgoals or many subgoals. Induction can produce a proof state that looks  | 
|
2359  | 
like this:  | 
|
2360  | 
\begin{isabelle}
 | 
|
2361  | 
\ 1.\ bigsubgoal1\isanewline  | 
|
2362  | 
\ 2.\ bigsubgoal2\isanewline  | 
|
2363  | 
\ 3.\ bigsubgoal3\isanewline  | 
|
2364  | 
\ 4.\ bigsubgoal4\isanewline  | 
|
2365  | 
\ 5.\ bigsubgoal5\isanewline  | 
|
2366  | 
\ 6.\ bigsubgoal6  | 
|
2367  | 
\end{isabelle}
 | 
|
2368  | 
If each \isa{bigsubgoal} is 15 lines or so, the proof state will be too big to
 | 
|
2369  | 
scroll through. By default, Isabelle displays at most 10 subgoals. The  | 
|
| 11406 | 2370  | 
\commdx{pr} command lets you change this limit:
 | 
| 10967 | 2371  | 
\begin{isabelle}
 | 
2372  | 
\isacommand{pr}\ 2\isanewline
 | 
|
2373  | 
\ 1.\ bigsubgoal1\isanewline  | 
|
2374  | 
\ 2.\ bigsubgoal2\isanewline  | 
|
2375  | 
A total of 6 subgoals...  | 
|
2376  | 
\end{isabelle}
 | 
|
2377  | 
||
2378  | 
\medskip  | 
|
2379  | 
All methods apply to the first subgoal.  | 
|
2380  | 
Sometimes, not only in a large proof, you may want to focus on some other  | 
|
2381  | 
subgoal.  Then you should try the commands \isacommand{defer} or \isacommand{prefer}.
 | 
|
2382  | 
||
2383  | 
In the following example, the first subgoal looks hard, while the others  | 
|
2384  | 
look as if \isa{blast} alone could prove them:
 | 
|
2385  | 
\begin{isabelle}
 | 
|
2386  | 
\ 1.\ hard\isanewline  | 
|
2387  | 
\ 2.\ \isasymnot \ \isasymnot \ P\ \isasymLongrightarrow \ P\isanewline  | 
|
2388  | 
\ 3.\ Q\ \isasymLongrightarrow \ Q%  | 
|
2389  | 
\end{isabelle}
 | 
|
2390  | 
%  | 
|
| 11406 | 2391  | 
The \commdx{defer} command moves the first subgoal into the last position.
 | 
| 10967 | 2392  | 
\begin{isabelle}
 | 
2393  | 
\isacommand{defer}\ 1\isanewline
 | 
|
2394  | 
\ 1.\ \isasymnot \ \isasymnot \ P\ \isasymLongrightarrow \ P\isanewline  | 
|
2395  | 
\ 2.\ Q\ \isasymLongrightarrow \ Q\isanewline  | 
|
2396  | 
\ 3.\ hard%  | 
|
2397  | 
\end{isabelle}
 | 
|
2398  | 
%  | 
|
2399  | 
Now we apply \isa{blast} repeatedly to the easy subgoals:
 | 
|
2400  | 
\begin{isabelle}
 | 
|
2401  | 
\isacommand{apply}\ blast+\isanewline
 | 
|
2402  | 
\ 1.\ hard%  | 
|
2403  | 
\end{isabelle}
 | 
|
2404  | 
Using \isacommand{defer}, we have cleared away the trivial parts of the proof so
 | 
|
2405  | 
that we can devote attention to the difficult part.  | 
|
2406  | 
||
2407  | 
\medskip  | 
|
| 11406 | 2408  | 
The \commdx{prefer} command moves the specified subgoal into the
 | 
| 10967 | 2409  | 
first position. For example, if you suspect that one of your subgoals is  | 
2410  | 
invalid (not a theorem), then you should investigate that subgoal first. If it  | 
|
2411  | 
cannot be proved, then there is no point in proving the other subgoals.  | 
|
2412  | 
\begin{isabelle}
 | 
|
2413  | 
\ 1.\ ok1\isanewline  | 
|
2414  | 
\ 2.\ ok2\isanewline  | 
|
2415  | 
\ 3.\ doubtful%  | 
|
2416  | 
\end{isabelle}
 | 
|
2417  | 
%  | 
|
2418  | 
We decide to work on the third subgoal.  | 
|
2419  | 
\begin{isabelle}
 | 
|
2420  | 
\isacommand{prefer}\ 3\isanewline
 | 
|
2421  | 
\ 1.\ doubtful\isanewline  | 
|
2422  | 
\ 2.\ ok1\isanewline  | 
|
2423  | 
\ 3.\ ok2  | 
|
2424  | 
\end{isabelle}
 | 
|
2425  | 
If we manage to prove \isa{doubtful}, then we can work on the other
 | 
|
2426  | 
subgoals, confident that we are not wasting our time. Finally we revise the  | 
|
2427  | 
proof script to remove the \isacommand{prefer} command, since we needed it only to
 | 
|
2428  | 
focus our exploration. The previous example is different: its use of  | 
|
2429  | 
\isacommand{defer} stops trivial subgoals from cluttering the rest of the
 | 
|
2430  | 
proof.  Even there, we should consider proving \isa{hard} as a preliminary
 | 
|
2431  | 
lemma. Always seek ways to streamline your proofs.  | 
|
2432  | 
||
2433  | 
||
2434  | 
\medskip  | 
|
2435  | 
Summary:  | 
|
2436  | 
\begin{itemize}
 | 
|
2437  | 
\item the control structures \isa+, \isa? and \isa| help express complicated proofs  | 
|
2438  | 
\item the \isacommand{pr} command can limit the number of subgoals to display
 | 
|
2439  | 
\item the \isacommand{defer} and \isacommand{prefer} commands move a 
 | 
|
2440  | 
subgoal to the last or first position  | 
|
2441  | 
\end{itemize}
 | 
|
2442  | 
||
2443  | 
\begin{exercise}
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2444  | 
Explain the use of \isa? and \isa+ in this proof.  | 
| 10967 | 2445  | 
\begin{isabelle}
 | 
2446  | 
\isacommand{lemma}\ "\isasymlbrakk P\isasymand Q\isasymlongrightarrow R;\ P\isasymlongrightarrow Q;\ P\isasymrbrakk \ \isasymLongrightarrow \ R"\isanewline
 | 
|
| 15617 | 2447  | 
\isacommand{by}\ (drule\ mp,\ (intro conjI)?,\ assumption+)+
 | 
| 10967 | 2448  | 
\end{isabelle}
 | 
2449  | 
\end{exercise}
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2450  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2451  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2452  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2453  | 
\section{Proving the Correctness of Euclid's Algorithm}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2454  | 
\label{sec:proving-euclid}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2455  | 
|
| 11406 | 2456  | 
\index{Euclid's algorithm|(}\index{*gcd (constant)|(}\index{divides relation|(}%
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2457  | 
A brief development will demonstrate the techniques of this chapter,  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2458  | 
including \isa{blast} applied with additional rules.  We shall also see
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2459  | 
\isa{case_tac} used to perform a Boolean case split.
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2460  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2461  | 
Let us prove that \isa{gcd} computes the greatest common
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2462  | 
divisor of its two arguments.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2463  | 
%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2464  | 
We use induction: \isa{gcd.induct} is the
 | 
| 25264 | 2465  | 
induction rule returned by \isa{fun}.  We simplify using
 | 
| 25258 | 2466  | 
rules proved in {\S}\ref{sec:fun-simplification}, since rewriting by the
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2467  | 
definition of \isa{gcd} can loop.
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2468  | 
\begin{isabelle}
 | 
| 25264 | 2469  | 
\isacommand{lemma}\ gcd\_dvd\_both:\ "(gcd\ m\ n\ dvd\ m)\ \isasymand \ (gcd\ m\ n\ dvd\ n)"
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2470  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2471  | 
The induction formula must be a conjunction. In the  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2472  | 
inductive step, each conjunct establishes the other.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2473  | 
\begin{isabelle}
 | 
| 25264 | 2474  | 
\ 1.\ \isasymAnd m\ n.\ (n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline  | 
2475  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ (}gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \isanewline
 | 
|
2476  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ (}gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n)\ \isasymLongrightarrow \isanewline
 | 
|
2477  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ }gcd\ m\ n\ dvd\ m\ \isasymand \ gcd\ m\ n\ dvd\ n%
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2478  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2479  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2480  | 
The conditional induction hypothesis suggests doing a case  | 
| 11406 | 2481  | 
analysis on \isa{n=0}.  We apply \methdx{case_tac} with type
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2482  | 
\isa{bool} --- and not with a datatype, as we have done until now.  Since
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2483  | 
\isa{nat} is a datatype, we could have written
 | 
| 12535 | 2484  | 
\isa{case_tac~n} instead of \isa{case_tac~"n=0"}.  However, the definition
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2485  | 
of \isa{gcd} makes a Boolean decision:
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2486  | 
\begin{isabelle}
 | 
| 25264 | 2487  | 
\ \ \ \ "gcd\ m\ n\ =\ (if\ n=0\ then\ m\ else\ gcd\ n\ (m\ mod\ n))"  | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2488  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2489  | 
Proofs about a function frequently follow the function's definition, so we perform  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2490  | 
case analysis over the same formula.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2491  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2492  | 
\isacommand{apply}\ (case_tac\ "n=0")\isanewline
 | 
| 25264 | 2493  | 
\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline  | 
2494  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk }gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \ gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n;\isanewline
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2495  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ \ }n\ =\ 0\isasymrbrakk \isanewline
 | 
| 25264 | 2496  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ gcd\ m\ n\ dvd\ m\ \isasymand \ gcd\ m\ n\ dvd\ n\isanewline
 | 
2497  | 
\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline  | 
|
2498  | 
\isaindent{\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk }gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \ gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n;\isanewline
 | 
|
2499  | 
\isaindent{\ 2.\ \isasymAnd m\ n.\ \ }n\ \isasymnoteq \ 0\isasymrbrakk \isanewline
 | 
|
2500  | 
\isaindent{\ 2.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ gcd\ m\ n\ dvd\ m\ \isasymand \ gcd\ m\ n\ dvd\ n%
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2501  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2502  | 
%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2503  | 
Simplification leaves one subgoal:  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2504  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2505  | 
\isacommand{apply}\ (simp_all)\isanewline
 | 
| 25264 | 2506  | 
\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \ gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n;\isanewline  | 
2507  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ \ }0\ <\ n\isasymrbrakk \isanewline
 | 
|
2508  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ gcd\ n\ (m\ mod\ n)\ dvd\ m%
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2509  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2510  | 
%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2511  | 
Here, we can use \isa{blast}.  
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2512  | 
One of the assumptions, the induction hypothesis, is a conjunction.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2513  | 
The two divides relationships it asserts are enough to prove  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2514  | 
the conclusion, for we have the following theorem at our disposal:  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2515  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2516  | 
\isasymlbrakk?k\ dvd\ (?m\ mod\ ?n){;}\ ?k\ dvd\ ?n\isasymrbrakk\ \isasymLongrightarrow\ ?k\ dvd\ ?m%
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2517  | 
\rulename{dvd_mod_imp_dvd}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2518  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2519  | 
%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2520  | 
This theorem can be applied in various ways. As an introduction rule, it  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2521  | 
would cause backward chaining from the conclusion (namely  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2522  | 
\isa{?k~dvd~?m}) to the two premises, which 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2523  | 
also involve the divides relation. This process does not look promising  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2524  | 
and could easily loop. More sensible is to apply the rule in the forward  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2525  | 
direction; each step would eliminate an occurrence of the \isa{mod} symbol, so the
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2526  | 
process must terminate.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2527  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2528  | 
\isacommand{apply}\ (blast\ dest:\ dvd_mod_imp_dvd)\isanewline
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2529  | 
\isacommand{done}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2530  | 
\end{isabelle}
 | 
| 11406 | 2531  | 
Attaching the \attrdx{dest} attribute to \isa{dvd_mod_imp_dvd} tells
 | 
2532  | 
\isa{blast} to use it as destruction rule; that is, in the forward direction.
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2533  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2534  | 
\medskip  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2535  | 
We have proved a conjunction. Now, let us give names to each of the  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2536  | 
two halves:  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2537  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2538  | 
\isacommand{lemmas}\ gcd_dvd1\ [iff]\ =\ gcd_dvd_both\ [THEN\ conjunct1]\isanewline
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2539  | 
\isacommand{lemmas}\ gcd_dvd2\ [iff]\ =\ gcd_dvd_both\ [THEN\ conjunct2]%
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2540  | 
\end{isabelle}
 | 
| 11406 | 2541  | 
Here we see \commdx{lemmas}
 | 
2542  | 
used with the \attrdx{iff} attribute, which supplies the new theorems to the
 | 
|
2543  | 
classical reasoner and the simplifier.  Recall that \attrdx{THEN} is
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2544  | 
frequently used with destruction rules; \isa{THEN conjunct1} extracts the
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2545  | 
first half of a conjunctive theorem.  Given \isa{gcd_dvd_both} it yields
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2546  | 
\begin{isabelle}
 | 
| 25264 | 2547  | 
\ \ \ \ \ gcd\ ?m1\ ?n1\ dvd\ ?m1  | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2548  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2549  | 
The variable names \isa{?m1} and \isa{?n1} arise because
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2550  | 
Isabelle renames schematic variables to prevent  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2551  | 
clashes.  The second \isacommand{lemmas} declaration yields
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2552  | 
\begin{isabelle}
 | 
| 25264 | 2553  | 
\ \ \ \ \ gcd\ ?m1\ ?n1\ dvd\ ?n1  | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2554  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2555  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2556  | 
\medskip  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2557  | 
To complete the verification of the \isa{gcd} function, we must 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2558  | 
prove that it returns the greatest of all the common divisors  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2559  | 
of its arguments. The proof is by induction, case analysis and simplification.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2560  | 
\begin{isabelle}
 | 
| 25264 | 2561  | 
\isacommand{lemma}\ gcd\_greatest\ [rule\_format]:\isanewline
 | 
2562  | 
\ \ \ \ \ \ "k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n"  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2563  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2564  | 
%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2565  | 
The goal is expressed using HOL implication,  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2566  | 
\isa{\isasymlongrightarrow}, because the induction affects the two
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2567  | 
preconditions.  The directive \isa{rule_format} tells Isabelle to replace
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2568  | 
each \isa{\isasymlongrightarrow} by \isa{\isasymLongrightarrow} before
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2569  | 
storing the eventual theorem. This directive can also remove outer  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2570  | 
universal quantifiers, converting the theorem into the usual format for  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2571  | 
inference rules. It can replace any series of applications of  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2572  | 
\isa{THEN} to the rules \isa{mp} and \isa{spec}.  We did not have to
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2573  | 
write this:  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2574  | 
\begin{isabelle}
 | 
| 25264 | 2575  | 
\isacommand{lemma}\ gcd_greatest\
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2576  | 
[THEN mp, THEN mp]:\isanewline  | 
| 25264 | 2577  | 
\ \ \ \ \ \ "k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n"  | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2578  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2579  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2580  | 
Because we are again reasoning about \isa{gcd}, we perform the same
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2581  | 
induction and case analysis as in the previous proof:  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2582  | 
\begingroup\samepage  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2583  | 
\begin{isabelle}
 | 
| 25264 | 2584  | 
\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline  | 
2585  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk }k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ m\ mod\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ n\ (m\ mod\ n);\isanewline
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2586  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ \ }n\ =\ 0\isasymrbrakk \isanewline
 | 
| 25264 | 2587  | 
\isaindent{\ 1.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n\isanewline
 | 
2588  | 
\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline  | 
|
2589  | 
\isaindent{\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk }k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ m\ mod\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ n\ (m\ mod\ n);\isanewline
 | 
|
2590  | 
\isaindent{\ 2.\ \isasymAnd m\ n.\ \ }n\ \isasymnoteq \ 0\isasymrbrakk \isanewline
 | 
|
2591  | 
\isaindent{\ 2.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n%
 | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2592  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2593  | 
\endgroup  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2594  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2595  | 
\noindent Simplification proves both subgoals.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2596  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2597  | 
\isacommand{apply}\ (simp_all\ add:\ dvd_mod)\isanewline
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2598  | 
\isacommand{done}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2599  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2600  | 
In the first, where \isa{n=0}, the implication becomes trivial: \isa{k\ dvd\
 | 
| 25264 | 2601  | 
gcd\ m\ n} goes to~\isa{k\ dvd\ m}.  The second subgoal is proved by
 | 
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2602  | 
an unfolding of \isa{gcd}, using this rule about divides:
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2603  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2604  | 
\isasymlbrakk ?f\ dvd\ ?m;\ ?f\ dvd\ ?n\isasymrbrakk \  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2605  | 
\isasymLongrightarrow \ ?f\ dvd\ ?m\ mod\ ?n%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2606  | 
\rulename{dvd_mod}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2607  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2608  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2609  | 
|
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2610  | 
\medskip  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2611  | 
The facts proved above can be summarized as a single logical  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2612  | 
equivalence. This step gives us a chance to see another application  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2613  | 
of \isa{blast}.
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2614  | 
\begin{isabelle}
 | 
| 25264 | 2615  | 
\isacommand{theorem}\ gcd\_greatest\_iff\ [iff]:\ \isanewline
 | 
2616  | 
\ \ \ \ \ \ \ \ "(k\ dvd\ gcd\ m\ n)\ =\ (k\ dvd\ m\ \isasymand \ k\ dvd\ n)"\isanewline  | 
|
| 
11080
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2617  | 
\isacommand{by}\ (blast\ intro!:\ gcd_greatest\ intro:\ dvd_trans)
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2618  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2619  | 
This theorem concisely expresses the correctness of the \isa{gcd} 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2620  | 
function.  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2621  | 
We state it with the \isa{iff} attribute so that 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2622  | 
Isabelle can use it to remove some occurrences of \isa{gcd}. 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2623  | 
The theorem has a one-line  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2624  | 
proof using \isa{blast} supplied with two additional introduction 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2625  | 
rules. The exclamation mark  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2626  | 
({\isa{intro}}{\isa{!}})\ signifies safe rules, which are 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2627  | 
applied aggressively. Rules given without the exclamation mark  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2628  | 
are applied reluctantly and their uses can be undone if  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2629  | 
the search backtracks. Here the unsafe rule expresses transitivity  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2630  | 
of the divides relation:  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2631  | 
\begin{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2632  | 
\isasymlbrakk?m\ dvd\ ?n;\ ?n\ dvd\ ?p\isasymrbrakk\ \isasymLongrightarrow\ ?m\ dvd\ ?p%  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2633  | 
\rulename{dvd_trans}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2634  | 
\end{isabelle}
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2635  | 
Applying \isa{dvd_trans} as 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2636  | 
an introduction rule entails a risk of looping, for it multiplies  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2637  | 
occurrences of the divides symbol. However, this proof relies  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2638  | 
on transitivity reasoning.  The rule {\isa{gcd\_greatest}} is safe to apply 
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2639  | 
aggressively because it yields simpler subgoals. The proof implicitly  | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2640  | 
uses \isa{gcd_dvd1} and \isa{gcd_dvd2} as safe rules, because they were
 | 
| 
 
22855d091249
various revisions in response to comments from Tobias
 
paulson 
parents: 
11077 
diff
changeset
 | 
2641  | 
declared using \isa{iff}.%
 | 
| 11406 | 2642  | 
\index{Euclid's algorithm|)}\index{*gcd (constant)|)}\index{divides relation|)}
 |