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