author | nipkow |
Tue, 29 Oct 2013 08:06:08 +0100 | |
changeset 54214 | 6d0688ce4ee2 |
parent 54212 | 5f1e2017eeea |
child 54217 | 2fa338fd0a28 |
permissions | -rw-r--r-- |
47269 | 1 |
(*<*) |
2 |
theory Logic |
|
3 |
imports LaTeXsugar |
|
4 |
begin |
|
5 |
(*>*) |
|
6 |
text{* |
|
7 |
\vspace{-5ex} |
|
51436 | 8 |
\section{Formulas} |
47269 | 9 |
|
47711 | 10 |
The core syntax of formulas (\textit{form} below) |
47720 | 11 |
provides the standard logical constructs, in decreasing order of precedence: |
47269 | 12 |
\[ |
13 |
\begin{array}{rcl} |
|
14 |
||
15 |
\mathit{form} & ::= & |
|
16 |
@{text"(form)"} ~\mid~ |
|
17 |
@{const True} ~\mid~ |
|
18 |
@{const False} ~\mid~ |
|
19 |
@{prop "term = term"}\\ |
|
20 |
&\mid& @{prop"\<not> form"} ~\mid~ |
|
21 |
@{prop "form \<and> form"} ~\mid~ |
|
22 |
@{prop "form \<or> form"} ~\mid~ |
|
23 |
@{prop "form \<longrightarrow> form"}\\ |
|
24 |
&\mid& @{prop"\<forall>x. form"} ~\mid~ @{prop"\<exists>x. form"} |
|
25 |
\end{array} |
|
26 |
\] |
|
47711 | 27 |
Terms are the ones we have seen all along, built from constants, variables, |
47269 | 28 |
function application and @{text"\<lambda>"}-abstraction, including all the syntactic |
29 |
sugar like infix symbols, @{text "if"}, @{text "case"} etc. |
|
30 |
\begin{warn} |
|
31 |
Remember that formulas are simply terms of type @{text bool}. Hence |
|
32 |
@{text "="} also works for formulas. Beware that @{text"="} has a higher |
|
33 |
precedence than the other logical operators. Hence @{prop"s = t \<and> A"} means |
|
47711 | 34 |
@{text"(s = t) \<and> A"}, and @{prop"A\<and>B = B\<and>A"} means @{text"A \<and> (B = B) \<and> A"}. |
47269 | 35 |
Logical equivalence can also be written with |
36 |
@{text "\<longleftrightarrow>"} instead of @{text"="}, where @{text"\<longleftrightarrow>"} has the same low |
|
37 |
precedence as @{text"\<longrightarrow>"}. Hence @{text"A \<and> B \<longleftrightarrow> B \<and> A"} really means |
|
38 |
@{text"(A \<and> B) \<longleftrightarrow> (B \<and> A)"}. |
|
39 |
\end{warn} |
|
40 |
\begin{warn} |
|
41 |
Quantifiers need to be enclosed in parentheses if they are nested within |
|
42 |
other constructs (just like @{text "if"}, @{text case} and @{text let}). |
|
43 |
\end{warn} |
|
52404 | 44 |
The most frequent logical symbols and their ASCII representations are shown |
45 |
in Fig.~\ref{fig:log-symbols}. |
|
46 |
\begin{figure} |
|
47269 | 47 |
\begin{center} |
48 |
\begin{tabular}{l@ {\qquad}l@ {\qquad}l} |
|
49 |
@{text "\<forall>"} & \xsymbol{forall} & \texttt{ALL}\\ |
|
50 |
@{text "\<exists>"} & \xsymbol{exists} & \texttt{EX}\\ |
|
51 |
@{text "\<lambda>"} & \xsymbol{lambda} & \texttt{\%}\\ |
|
52 |
@{text "\<longrightarrow>"} & \texttt{-{}->}\\ |
|
47711 | 53 |
@{text "\<longleftrightarrow>"} & \texttt{<->}\\ |
47269 | 54 |
@{text "\<and>"} & \texttt{/\char`\\} & \texttt{\&}\\ |
55 |
@{text "\<or>"} & \texttt{\char`\\/} & \texttt{|}\\ |
|
56 |
@{text "\<not>"} & \xsymbol{not} & \texttt{\char`~}\\ |
|
57 |
@{text "\<noteq>"} & \xsymbol{noteq} & \texttt{\char`~=} |
|
58 |
\end{tabular} |
|
59 |
\end{center} |
|
52404 | 60 |
\caption{Logical symbols and their ASCII forms} |
61 |
\label{fig:log-symbols} |
|
62 |
\end{figure} |
|
63 |
The first column shows the symbols, the other columns ASCII representations. |
|
64 |
The \texttt{\char`\\}\texttt{<...>} form is always converted into the symbolic form |
|
65 |
by the Isabelle interfaces, the treatment of the other ASCII forms |
|
66 |
depends on the interface. The ASCII forms \texttt{/\char`\\} and |
|
67 |
\texttt{\char`\\/} |
|
68 |
are special in that they are merely keyboard shortcuts for the interface and |
|
69 |
not logical symbols by themselves. |
|
47269 | 70 |
\begin{warn} |
71 |
The implication @{text"\<Longrightarrow>"} is part of the Isabelle framework. It structures |
|
47711 | 72 |
theorems and proof states, separating assumptions from conclusions. |
47269 | 73 |
The implication @{text"\<longrightarrow>"} is part of the logic HOL and can occur inside the |
74 |
formulas that make up the assumptions and conclusion. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
75 |
Theorems should be of the form @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"}, |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
76 |
not @{text"A\<^sub>1 \<and> \<dots> \<and> A\<^sub>n \<longrightarrow> A"}. Both are logically equivalent |
47269 | 77 |
but the first one works better when using the theorem in further proofs. |
78 |
\end{warn} |
|
79 |
||
51436 | 80 |
\section{Sets} |
51038 | 81 |
\label{sec:Sets} |
47269 | 82 |
|
83 |
Sets of elements of type @{typ 'a} have type @{typ"'a set"}. |
|
47711 | 84 |
They can be finite or infinite. Sets come with the usual notation: |
47269 | 85 |
\begin{itemize} |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
86 |
\item @{term"{}"},\quad @{text"{e\<^sub>1,\<dots>,e\<^sub>n}"} |
47269 | 87 |
\item @{prop"e \<in> A"},\quad @{prop"A \<subseteq> B"} |
88 |
\item @{term"A \<union> B"},\quad @{term"A \<inter> B"},\quad @{term"A - B"},\quad @{term"-A"} |
|
89 |
\end{itemize} |
|
51425 | 90 |
(where @{term"A-B"} and @{text"-A"} are set difference and complement) |
47269 | 91 |
and much more. @{const UNIV} is the set of all elements of some type. |
92 |
Set comprehension is written @{term"{x. P}"} |
|
51425 | 93 |
rather than @{text"{x | P}"}. |
47269 | 94 |
\begin{warn} |
95 |
In @{term"{x. P}"} the @{text x} must be a variable. Set comprehension |
|
96 |
involving a proper term @{text t} must be written |
|
49615 | 97 |
\noquotes{@{term[source] "{t | x y. P}"}}, |
98 |
where @{text "x y"} are those free variables in @{text t} |
|
99 |
that occur in @{text P}. |
|
100 |
This is just a shorthand for @{term"{v. EX x y. v = t \<and> P}"}, where |
|
101 |
@{text v} is a new variable. For example, @{term"{x+y|x. x \<in> A}"} |
|
102 |
is short for \noquotes{@{term[source]"{v. \<exists>x. v = x+y \<and> x \<in> A}"}}. |
|
47269 | 103 |
\end{warn} |
104 |
||
105 |
Here are the ASCII representations of the mathematical symbols: |
|
106 |
\begin{center} |
|
107 |
\begin{tabular}{l@ {\quad}l@ {\quad}l} |
|
108 |
@{text "\<in>"} & \texttt{\char`\\\char`\<in>} & \texttt{:}\\ |
|
109 |
@{text "\<subseteq>"} & \texttt{\char`\\\char`\<subseteq>} & \texttt{<=}\\ |
|
110 |
@{text "\<union>"} & \texttt{\char`\\\char`\<union>} & \texttt{Un}\\ |
|
111 |
@{text "\<inter>"} & \texttt{\char`\\\char`\<inter>} & \texttt{Int} |
|
112 |
\end{tabular} |
|
113 |
\end{center} |
|
114 |
Sets also allow bounded quantifications @{prop"ALL x : A. P"} and |
|
115 |
@{prop"EX x : A. P"}. |
|
116 |
||
52404 | 117 |
For the more ambitious, there are also @{text"\<Union>"} and @{text"\<Inter>"}: |
118 |
\begin{center} |
|
119 |
@{thm Union_eq} \qquad @{thm Inter_eq} |
|
120 |
\end{center} |
|
121 |
The ASCII forms of @{text"\<Union>"} are \texttt{\char`\\\char`\<Union>} and \texttt{Union}, |
|
122 |
those of @{text"\<Inter>"} are \texttt{\char`\\\char`\<Inter>} and \texttt{Inter}. |
|
123 |
There are also indexed unions and intersections: |
|
51784 | 124 |
\begin{center} |
52404 | 125 |
@{thm UNION_eq} \\ @{thm INTER_eq} |
126 |
\end{center} |
|
127 |
The ASCII forms are \ \texttt{UN x:A.~B} \ and \ \texttt{INT x:A. B} \ |
|
128 |
where \texttt{x} may occur in \texttt{B}. |
|
129 |
If \texttt{A} is \texttt{UNIV} you can just write \ \texttt{UN x.~B} \ and \ \texttt{INT x. B}. |
|
130 |
||
131 |
Some other frequently useful functions on sets are the following: |
|
132 |
\begin{center} |
|
133 |
\begin{tabular}{l@ {\quad}l} |
|
51784 | 134 |
@{const_typ set} & converts a list to the set of its elements\\ |
135 |
@{const_typ finite} & is true iff its argument is finite\\ |
|
136 |
@{const_typ card} & is the cardinality of a finite set\\ |
|
137 |
& and is @{text 0} for all infinite sets\\ |
|
138 |
@{thm image_def} & is the image of a function over a set |
|
139 |
\end{tabular} |
|
140 |
\end{center} |
|
141 |
See \cite{Nipkow-Main} for the wealth of further predefined functions in theory |
|
142 |
@{theory Main}. |
|
143 |
||
54214 | 144 |
|
145 |
\subsection{Exercises} |
|
146 |
||
147 |
\exercise |
|
148 |
Start from the data type of binary trees defined earlier: |
|
149 |
*} |
|
150 |
||
151 |
datatype 'a tree = Tip | Node "'a tree" 'a "'a tree" |
|
152 |
||
153 |
text{* |
|
154 |
Define a function @{text "set ::"} @{typ "'a tree \<Rightarrow> 'a set"} |
|
155 |
that returns the elements in a tree and a function |
|
156 |
@{text "ord ::"} @{typ "int tree \<Rightarrow> bool"} |
|
157 |
the tests if an @{typ "int tree"} is ordered. |
|
158 |
||
159 |
Define a function @{text ins} that inserts an element into an ordered @{typ "int tree"} |
|
160 |
while maintaining the order of the tree. If the element is already in the tree, the |
|
161 |
same tree should be returned. Prove correctness of @{text ins}: |
|
162 |
@{prop "set(ins x t) = {x} \<union> set t"} and @{prop "ord t \<Longrightarrow> ord(ins i t)"}. |
|
163 |
\endexercise |
|
164 |
||
165 |
||
52361 | 166 |
\section{Proof Automation} |
47269 | 167 |
|
168 |
So far we have only seen @{text simp} and @{text auto}: Both perform |
|
169 |
rewriting, both can also prove linear arithmetic facts (no multiplication), |
|
170 |
and @{text auto} is also able to prove simple logical or set-theoretic goals: |
|
171 |
*} |
|
172 |
||
173 |
lemma "\<forall>x. \<exists>y. x = y" |
|
174 |
by auto |
|
175 |
||
176 |
lemma "A \<subseteq> B \<inter> C \<Longrightarrow> A \<subseteq> B \<union> C" |
|
177 |
by auto |
|
178 |
||
179 |
text{* where |
|
180 |
\begin{quote} |
|
181 |
\isacom{by} \textit{proof-method} |
|
182 |
\end{quote} |
|
183 |
is short for |
|
184 |
\begin{quote} |
|
185 |
\isacom{apply} \textit{proof-method}\\ |
|
186 |
\isacom{done} |
|
187 |
\end{quote} |
|
188 |
The key characteristics of both @{text simp} and @{text auto} are |
|
189 |
\begin{itemize} |
|
190 |
\item They show you were they got stuck, giving you an idea how to continue. |
|
191 |
\item They perform the obvious steps but are highly incomplete. |
|
192 |
\end{itemize} |
|
193 |
A proof method is \concept{complete} if it can prove all true formulas. |
|
194 |
There is no complete proof method for HOL, not even in theory. |
|
195 |
Hence all our proof methods only differ in how incomplete they are. |
|
196 |
||
197 |
A proof method that is still incomplete but tries harder than @{text auto} is |
|
198 |
@{text fastforce}. It either succeeds or fails, it acts on the first |
|
199 |
subgoal only, and it can be modified just like @{text auto}, e.g.\ |
|
200 |
with @{text "simp add"}. Here is a typical example of what @{text fastforce} |
|
201 |
can do: |
|
202 |
*} |
|
203 |
||
204 |
lemma "\<lbrakk> \<forall>xs \<in> A. \<exists>ys. xs = ys @ ys; us \<in> A \<rbrakk> |
|
205 |
\<Longrightarrow> \<exists>n. length us = n+n" |
|
206 |
by fastforce |
|
207 |
||
208 |
text{* This lemma is out of reach for @{text auto} because of the |
|
209 |
quantifiers. Even @{text fastforce} fails when the quantifier structure |
|
210 |
becomes more complicated. In a few cases, its slow version @{text force} |
|
211 |
succeeds where @{text fastforce} fails. |
|
212 |
||
213 |
The method of choice for complex logical goals is @{text blast}. In the |
|
47711 | 214 |
following example, @{text T} and @{text A} are two binary predicates. It |
215 |
is shown that if @{text T} is total, @{text A} is antisymmetric and @{text T} is |
|
47269 | 216 |
a subset of @{text A}, then @{text A} is a subset of @{text T}: |
217 |
*} |
|
218 |
||
219 |
lemma |
|
220 |
"\<lbrakk> \<forall>x y. T x y \<or> T y x; |
|
221 |
\<forall>x y. A x y \<and> A y x \<longrightarrow> x = y; |
|
222 |
\<forall>x y. T x y \<longrightarrow> A x y \<rbrakk> |
|
223 |
\<Longrightarrow> \<forall>x y. A x y \<longrightarrow> T x y" |
|
224 |
by blast |
|
225 |
||
226 |
text{* |
|
227 |
We leave it to the reader to figure out why this lemma is true. |
|
228 |
Method @{text blast} |
|
229 |
\begin{itemize} |
|
230 |
\item is (in principle) a complete proof procedure for first-order formulas, |
|
231 |
a fragment of HOL. In practice there is a search bound. |
|
232 |
\item does no rewriting and knows very little about equality. |
|
233 |
\item covers logic, sets and relations. |
|
234 |
\item either succeeds or fails. |
|
235 |
\end{itemize} |
|
236 |
Because of its strength in logic and sets and its weakness in equality reasoning, it complements the earlier proof methods. |
|
237 |
||
238 |
||
51436 | 239 |
\subsection{Sledgehammer} |
47269 | 240 |
|
241 |
Command \isacom{sledgehammer} calls a number of external automatic |
|
242 |
theorem provers (ATPs) that run for up to 30 seconds searching for a |
|
243 |
proof. Some of these ATPs are part of the Isabelle installation, others are |
|
244 |
queried over the internet. If successful, a proof command is generated and can |
|
245 |
be inserted into your proof. The biggest win of \isacom{sledgehammer} is |
|
246 |
that it will take into account the whole lemma library and you do not need to |
|
247 |
feed in any lemma explicitly. For example,*} |
|
248 |
||
249 |
lemma "\<lbrakk> xs @ ys = ys @ xs; length xs = length ys \<rbrakk> \<Longrightarrow> xs = ys" |
|
250 |
||
251 |
txt{* cannot be solved by any of the standard proof methods, but |
|
252 |
\isacom{sledgehammer} finds the following proof: *} |
|
253 |
||
254 |
by (metis append_eq_conv_conj) |
|
255 |
||
256 |
text{* We do not explain how the proof was found but what this command |
|
257 |
means. For a start, Isabelle does not trust external tools (and in particular |
|
258 |
not the translations from Isabelle's logic to those tools!) |
|
259 |
and insists on a proof that it can check. This is what @{text metis} does. |
|
260 |
It is given a list of lemmas and tries to find a proof just using those lemmas |
|
261 |
(and pure logic). In contrast to @{text simp} and friends that know a lot of |
|
262 |
lemmas already, using @{text metis} manually is tedious because one has |
|
263 |
to find all the relevant lemmas first. But that is precisely what |
|
264 |
\isacom{sledgehammer} does for us. |
|
265 |
In this case lemma @{thm[source]append_eq_conv_conj} alone suffices: |
|
266 |
@{thm[display] append_eq_conv_conj} |
|
267 |
We leave it to the reader to figure out why this lemma suffices to prove |
|
268 |
the above lemma, even without any knowledge of what the functions @{const take} |
|
269 |
and @{const drop} do. Keep in mind that the variables in the two lemmas |
|
270 |
are independent of each other, despite the same names, and that you can |
|
271 |
substitute arbitrary values for the free variables in a lemma. |
|
272 |
||
273 |
Just as for the other proof methods we have seen, there is no guarantee that |
|
274 |
\isacom{sledgehammer} will find a proof if it exists. Nor is |
|
275 |
\isacom{sledgehammer} superior to the other proof methods. They are |
|
276 |
incomparable. Therefore it is recommended to apply @{text simp} or @{text |
|
277 |
auto} before invoking \isacom{sledgehammer} on what is left. |
|
278 |
||
51436 | 279 |
\subsection{Arithmetic} |
47269 | 280 |
|
281 |
By arithmetic formulas we mean formulas involving variables, numbers, @{text |
|
282 |
"+"}, @{text"-"}, @{text "="}, @{text "<"}, @{text "\<le>"} and the usual logical |
|
283 |
connectives @{text"\<not>"}, @{text"\<and>"}, @{text"\<or>"}, @{text"\<longrightarrow>"}, |
|
284 |
@{text"\<longleftrightarrow>"}. Strictly speaking, this is known as \concept{linear arithmetic} |
|
285 |
because it does not involve multiplication, although multiplication with |
|
286 |
numbers, e.g.\ @{text"2*n"} is allowed. Such formulas can be proved by |
|
287 |
@{text arith}: |
|
288 |
*} |
|
289 |
||
290 |
lemma "\<lbrakk> (a::nat) \<le> x + b; 2*x < c \<rbrakk> \<Longrightarrow> 2*a + 1 \<le> 2*b + c" |
|
291 |
by arith |
|
292 |
||
293 |
text{* In fact, @{text auto} and @{text simp} can prove many linear |
|
294 |
arithmetic formulas already, like the one above, by calling a weak but fast |
|
295 |
version of @{text arith}. Hence it is usually not necessary to invoke |
|
296 |
@{text arith} explicitly. |
|
297 |
||
298 |
The above example involves natural numbers, but integers (type @{typ int}) |
|
299 |
and real numbers (type @{text real}) are supported as well. As are a number |
|
300 |
of further operators like @{const min} and @{const max}. On @{typ nat} and |
|
301 |
@{typ int}, @{text arith} can even prove theorems with quantifiers in them, |
|
302 |
but we will not enlarge on that here. |
|
303 |
||
304 |
||
52361 | 305 |
\subsection{Trying Them All} |
47727 | 306 |
|
307 |
If you want to try all of the above automatic proof methods you simply type |
|
308 |
\begin{isabelle} |
|
309 |
\isacom{try} |
|
310 |
\end{isabelle} |
|
311 |
You can also add specific simplification and introduction rules: |
|
312 |
\begin{isabelle} |
|
313 |
\isacom{try} @{text"simp: \<dots> intro: \<dots>"} |
|
314 |
\end{isabelle} |
|
315 |
There is also a lightweight variant \isacom{try0} that does not call |
|
316 |
sledgehammer. |
|
317 |
||
52361 | 318 |
\section{Single Step Proofs} |
47269 | 319 |
|
320 |
Although automation is nice, it often fails, at least initially, and you need |
|
321 |
to find out why. When @{text fastforce} or @{text blast} simply fail, you have |
|
322 |
no clue why. At this point, the stepwise |
|
323 |
application of proof rules may be necessary. For example, if @{text blast} |
|
324 |
fails on @{prop"A \<and> B"}, you want to attack the two |
|
325 |
conjuncts @{text A} and @{text B} separately. This can |
|
326 |
be achieved by applying \emph{conjunction introduction} |
|
327 |
\[ @{thm[mode=Rule,show_question_marks]conjI}\ @{text conjI} |
|
328 |
\] |
|
329 |
to the proof state. We will now examine the details of this process. |
|
330 |
||
52361 | 331 |
\subsection{Instantiating Unknowns} |
47269 | 332 |
|
333 |
We had briefly mentioned earlier that after proving some theorem, |
|
334 |
Isabelle replaces all free variables @{text x} by so called \concept{unknowns} |
|
335 |
@{text "?x"}. We can see this clearly in rule @{thm[source] conjI}. |
|
336 |
These unknowns can later be instantiated explicitly or implicitly: |
|
337 |
\begin{itemize} |
|
338 |
\item By hand, using @{text of}. |
|
339 |
The expression @{text"conjI[of \"a=b\" \"False\"]"} |
|
340 |
instantiates the unknowns in @{thm[source] conjI} from left to right with the |
|
341 |
two formulas @{text"a=b"} and @{text False}, yielding the rule |
|
342 |
@{thm[display,mode=Rule]conjI[of "a=b" False]} |
|
343 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
344 |
In general, @{text"th[of string\<^sub>1 \<dots> string\<^sub>n]"} instantiates |
47269 | 345 |
the unknowns in the theorem @{text th} from left to right with the terms |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
346 |
@{text string\<^sub>1} to @{text string\<^sub>n}. |
47269 | 347 |
|
348 |
\item By unification. \concept{Unification} is the process of making two |
|
349 |
terms syntactically equal by suitable instantiations of unknowns. For example, |
|
350 |
unifying @{text"?P \<and> ?Q"} with \mbox{@{prop"a=b \<and> False"}} instantiates |
|
351 |
@{text "?P"} with @{prop "a=b"} and @{text "?Q"} with @{prop False}. |
|
352 |
\end{itemize} |
|
353 |
We need not instantiate all unknowns. If we want to skip a particular one we |
|
354 |
can just write @{text"_"} instead, for example @{text "conjI[of _ \"False\"]"}. |
|
355 |
Unknowns can also be instantiated by name, for example |
|
356 |
@{text "conjI[where ?P = \"a=b\" and ?Q = \"False\"]"}. |
|
357 |
||
358 |
||
52361 | 359 |
\subsection{Rule Application} |
47269 | 360 |
|
361 |
\concept{Rule application} means applying a rule backwards to a proof state. |
|
362 |
For example, applying rule @{thm[source]conjI} to a proof state |
|
363 |
\begin{quote} |
|
364 |
@{text"1. \<dots> \<Longrightarrow> A \<and> B"} |
|
365 |
\end{quote} |
|
366 |
results in two subgoals, one for each premise of @{thm[source]conjI}: |
|
367 |
\begin{quote} |
|
368 |
@{text"1. \<dots> \<Longrightarrow> A"}\\ |
|
369 |
@{text"2. \<dots> \<Longrightarrow> B"} |
|
370 |
\end{quote} |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
371 |
In general, the application of a rule @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"} |
47269 | 372 |
to a subgoal \mbox{@{text"\<dots> \<Longrightarrow> C"}} proceeds in two steps: |
373 |
\begin{enumerate} |
|
374 |
\item |
|
375 |
Unify @{text A} and @{text C}, thus instantiating the unknowns in the rule. |
|
376 |
\item |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
377 |
Replace the subgoal @{text C} with @{text n} new subgoals @{text"A\<^sub>1"} to @{text"A\<^sub>n"}. |
47269 | 378 |
\end{enumerate} |
379 |
This is the command to apply rule @{text xyz}: |
|
380 |
\begin{quote} |
|
381 |
\isacom{apply}@{text"(rule xyz)"} |
|
382 |
\end{quote} |
|
383 |
This is also called \concept{backchaining} with rule @{text xyz}. |
|
384 |
||
52361 | 385 |
\subsection{Introduction Rules} |
47269 | 386 |
|
387 |
Conjunction introduction (@{thm[source] conjI}) is one example of a whole |
|
388 |
class of rules known as \concept{introduction rules}. They explain under which |
|
389 |
premises some logical construct can be introduced. Here are some further |
|
390 |
useful introduction rules: |
|
391 |
\[ |
|
392 |
\inferrule*[right=\mbox{@{text impI}}]{\mbox{@{text"?P \<Longrightarrow> ?Q"}}}{\mbox{@{text"?P \<longrightarrow> ?Q"}}} |
|
393 |
\qquad |
|
394 |
\inferrule*[right=\mbox{@{text allI}}]{\mbox{@{text"\<And>x. ?P x"}}}{\mbox{@{text"\<forall>x. ?P x"}}} |
|
395 |
\] |
|
396 |
\[ |
|
397 |
\inferrule*[right=\mbox{@{text iffI}}]{\mbox{@{text"?P \<Longrightarrow> ?Q"}} \\ \mbox{@{text"?Q \<Longrightarrow> ?P"}}} |
|
398 |
{\mbox{@{text"?P = ?Q"}}} |
|
399 |
\] |
|
400 |
These rules are part of the logical system of \concept{natural deduction} |
|
401 |
(e.g.\ \cite{HuthRyan}). Although we intentionally de-emphasize the basic rules |
|
402 |
of logic in favour of automatic proof methods that allow you to take bigger |
|
403 |
steps, these rules are helpful in locating where and why automation fails. |
|
404 |
When applied backwards, these rules decompose the goal: |
|
405 |
\begin{itemize} |
|
406 |
\item @{thm[source] conjI} and @{thm[source]iffI} split the goal into two subgoals, |
|
407 |
\item @{thm[source] impI} moves the left-hand side of a HOL implication into the list of assumptions, |
|
408 |
\item and @{thm[source] allI} removes a @{text "\<forall>"} by turning the quantified variable into a fixed local variable of the subgoal. |
|
409 |
\end{itemize} |
|
410 |
Isabelle knows about these and a number of other introduction rules. |
|
411 |
The command |
|
412 |
\begin{quote} |
|
413 |
\isacom{apply} @{text rule} |
|
414 |
\end{quote} |
|
415 |
automatically selects the appropriate rule for the current subgoal. |
|
416 |
||
417 |
You can also turn your own theorems into introduction rules by giving them |
|
47711 | 418 |
the @{text"intro"} attribute, analogous to the @{text simp} attribute. In |
47269 | 419 |
that case @{text blast}, @{text fastforce} and (to a limited extent) @{text |
420 |
auto} will automatically backchain with those theorems. The @{text intro} |
|
421 |
attribute should be used with care because it increases the search space and |
|
47711 | 422 |
can lead to nontermination. Sometimes it is better to use it only in |
423 |
specific calls of @{text blast} and friends. For example, |
|
47269 | 424 |
@{thm[source] le_trans}, transitivity of @{text"\<le>"} on type @{typ nat}, |
425 |
is not an introduction rule by default because of the disastrous effect |
|
426 |
on the search space, but can be useful in specific situations: |
|
427 |
*} |
|
428 |
||
429 |
lemma "\<lbrakk> (a::nat) \<le> b; b \<le> c; c \<le> d; d \<le> e \<rbrakk> \<Longrightarrow> a \<le> e" |
|
430 |
by(blast intro: le_trans) |
|
431 |
||
432 |
text{* |
|
433 |
Of course this is just an example and could be proved by @{text arith}, too. |
|
434 |
||
52361 | 435 |
\subsection{Forward Proof} |
47269 | 436 |
\label{sec:forward-proof} |
437 |
||
438 |
Forward proof means deriving new theorems from old theorems. We have already |
|
439 |
seen a very simple form of forward proof: the @{text of} operator for |
|
440 |
instantiating unknowns in a theorem. The big brother of @{text of} is @{text |
|
441 |
OF} for applying one theorem to others. Given a theorem @{prop"A \<Longrightarrow> B"} called |
|
442 |
@{text r} and a theorem @{text A'} called @{text r'}, the theorem @{text |
|
443 |
"r[OF r']"} is the result of applying @{text r} to @{text r'}, where @{text |
|
444 |
r} should be viewed as a function taking a theorem @{text A} and returning |
|
445 |
@{text B}. More precisely, @{text A} and @{text A'} are unified, thus |
|
446 |
instantiating the unknowns in @{text B}, and the result is the instantiated |
|
447 |
@{text B}. Of course, unification may also fail. |
|
448 |
\begin{warn} |
|
449 |
Application of rules to other rules operates in the forward direction: from |
|
450 |
the premises to the conclusion of the rule; application of rules to proof |
|
451 |
states operates in the backward direction, from the conclusion to the |
|
452 |
premises. |
|
453 |
\end{warn} |
|
454 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
455 |
In general @{text r} can be of the form @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"} |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
456 |
and there can be multiple argument theorems @{text r\<^sub>1} to @{text r\<^sub>m} |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
457 |
(with @{text"m \<le> n"}), in which case @{text "r[OF r\<^sub>1 \<dots> r\<^sub>m]"} is obtained |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
458 |
by unifying and thus proving @{text "A\<^sub>i"} with @{text "r\<^sub>i"}, @{text"i = 1\<dots>m"}. |
47269 | 459 |
Here is an example, where @{thm[source]refl} is the theorem |
460 |
@{thm[show_question_marks] refl}: |
|
461 |
*} |
|
462 |
||
463 |
thm conjI[OF refl[of "a"] refl[of "b"]] |
|
464 |
||
465 |
text{* yields the theorem @{thm conjI[OF refl[of "a"] refl[of "b"]]}. |
|
466 |
The command \isacom{thm} merely displays the result. |
|
467 |
||
468 |
Forward reasoning also makes sense in connection with proof states. |
|
469 |
Therefore @{text blast}, @{text fastforce} and @{text auto} support a modifier |
|
470 |
@{text dest} which instructs the proof method to use certain rules in a |
|
471 |
forward fashion. If @{text r} is of the form \mbox{@{text "A \<Longrightarrow> B"}}, the modifier |
|
472 |
\mbox{@{text"dest: r"}} |
|
473 |
allows proof search to reason forward with @{text r}, i.e.\ |
|
474 |
to replace an assumption @{text A'}, where @{text A'} unifies with @{text A}, |
|
475 |
with the correspondingly instantiated @{text B}. For example, @{thm[source,show_question_marks] Suc_leD} is the theorem \mbox{@{thm Suc_leD}}, which works well for forward reasoning: |
|
476 |
*} |
|
477 |
||
478 |
lemma "Suc(Suc(Suc a)) \<le> b \<Longrightarrow> a \<le> b" |
|
479 |
by(blast dest: Suc_leD) |
|
480 |
||
481 |
text{* In this particular example we could have backchained with |
|
482 |
@{thm[source] Suc_leD}, too, but because the premise is more complicated than the conclusion this can easily lead to nontermination. |
|
483 |
||
54212 | 484 |
%\subsection{Finding Theorems} |
485 |
% |
|
486 |
%Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current |
|
487 |
%theory. Search criteria include pattern matching on terms and on names. |
|
488 |
%For details see the Isabelle/Isar Reference Manual~\cite{IsarRef}. |
|
489 |
%\bigskip |
|
47727 | 490 |
|
47269 | 491 |
\begin{warn} |
492 |
To ease readability we will drop the question marks |
|
493 |
in front of unknowns from now on. |
|
494 |
\end{warn} |
|
495 |
||
47727 | 496 |
|
52361 | 497 |
\section{Inductive Definitions} |
47269 | 498 |
\label{sec:inductive-defs} |
499 |
||
500 |
Inductive definitions are the third important definition facility, after |
|
47711 | 501 |
datatypes and recursive function. |
52782 | 502 |
\ifsem |
47711 | 503 |
In fact, they are the key construct in the |
47269 | 504 |
definition of operational semantics in the second part of the book. |
52782 | 505 |
\fi |
47269 | 506 |
|
52361 | 507 |
\subsection{An Example: Even Numbers} |
47269 | 508 |
\label{sec:Logic:even} |
509 |
||
510 |
Here is a simple example of an inductively defined predicate: |
|
511 |
\begin{itemize} |
|
512 |
\item 0 is even |
|
513 |
\item If $n$ is even, so is $n+2$. |
|
514 |
\end{itemize} |
|
515 |
The operative word ``inductive'' means that these are the only even numbers. |
|
516 |
In Isabelle we give the two rules the names @{text ev0} and @{text evSS} |
|
517 |
and write |
|
518 |
*} |
|
519 |
||
520 |
inductive ev :: "nat \<Rightarrow> bool" where |
|
521 |
ev0: "ev 0" | |
|
522 |
evSS: (*<*)"ev n \<Longrightarrow> ev (Suc(Suc n))"(*>*) |
|
47306 | 523 |
text_raw{* @{prop[source]"ev n \<Longrightarrow> ev (n + 2)"} *} |
47269 | 524 |
|
525 |
text{* To get used to inductive definitions, we will first prove a few |
|
526 |
properties of @{const ev} informally before we descend to the Isabelle level. |
|
527 |
||
528 |
How do we prove that some number is even, e.g.\ @{prop "ev 4"}? Simply by combining the defining rules for @{const ev}: |
|
529 |
\begin{quote} |
|
530 |
@{text "ev 0 \<Longrightarrow> ev (0 + 2) \<Longrightarrow> ev((0 + 2) + 2) = ev 4"} |
|
531 |
\end{quote} |
|
532 |
||
52361 | 533 |
\subsubsection{Rule Induction} |
47269 | 534 |
|
535 |
Showing that all even numbers have some property is more complicated. For |
|
536 |
example, let us prove that the inductive definition of even numbers agrees |
|
537 |
with the following recursive one:*} |
|
538 |
||
539 |
fun even :: "nat \<Rightarrow> bool" where |
|
540 |
"even 0 = True" | |
|
541 |
"even (Suc 0) = False" | |
|
542 |
"even (Suc(Suc n)) = even n" |
|
543 |
||
544 |
text{* We prove @{prop"ev m \<Longrightarrow> even m"}. That is, we |
|
545 |
assume @{prop"ev m"} and by induction on the form of its derivation |
|
546 |
prove @{prop"even m"}. There are two cases corresponding to the two rules |
|
547 |
for @{const ev}: |
|
548 |
\begin{description} |
|
549 |
\item[Case @{thm[source]ev0}:] |
|
550 |
@{prop"ev m"} was derived by rule @{prop "ev 0"}: \\ |
|
551 |
@{text"\<Longrightarrow>"} @{prop"m=(0::nat)"} @{text"\<Longrightarrow>"} @{text "even m = even 0 = True"} |
|
552 |
\item[Case @{thm[source]evSS}:] |
|
553 |
@{prop"ev m"} was derived by rule @{prop "ev n \<Longrightarrow> ev(n+2)"}: \\ |
|
554 |
@{text"\<Longrightarrow>"} @{prop"m=n+(2::nat)"} and by induction hypothesis @{prop"even n"}\\ |
|
555 |
@{text"\<Longrightarrow>"} @{text"even m = even(n + 2) = even n = True"} |
|
556 |
\end{description} |
|
557 |
||
558 |
What we have just seen is a special case of \concept{rule induction}. |
|
559 |
Rule induction applies to propositions of this form |
|
560 |
\begin{quote} |
|
561 |
@{prop "ev n \<Longrightarrow> P n"} |
|
562 |
\end{quote} |
|
563 |
That is, we want to prove a property @{prop"P n"} |
|
564 |
for all even @{text n}. But if we assume @{prop"ev n"}, then there must be |
|
565 |
some derivation of this assumption using the two defining rules for |
|
566 |
@{const ev}. That is, we must prove |
|
567 |
\begin{description} |
|
568 |
\item[Case @{thm[source]ev0}:] @{prop"P(0::nat)"} |
|
569 |
\item[Case @{thm[source]evSS}:] @{prop"\<lbrakk> ev n; P n \<rbrakk> \<Longrightarrow> P(n + 2::nat)"} |
|
570 |
\end{description} |
|
571 |
The corresponding rule is called @{thm[source] ev.induct} and looks like this: |
|
572 |
\[ |
|
573 |
\inferrule{ |
|
574 |
\mbox{@{thm (prem 1) ev.induct[of "n"]}}\\ |
|
575 |
\mbox{@{thm (prem 2) ev.induct}}\\ |
|
576 |
\mbox{@{prop"!!n. \<lbrakk> ev n; P n \<rbrakk> \<Longrightarrow> P(n+2)"}}} |
|
577 |
{\mbox{@{thm (concl) ev.induct[of "n"]}}} |
|
578 |
\] |
|
579 |
The first premise @{prop"ev n"} enforces that this rule can only be applied |
|
580 |
in situations where we know that @{text n} is even. |
|
581 |
||
582 |
Note that in the induction step we may not just assume @{prop"P n"} but also |
|
583 |
\mbox{@{prop"ev n"}}, which is simply the premise of rule @{thm[source] |
|
584 |
evSS}. Here is an example where the local assumption @{prop"ev n"} comes in |
|
585 |
handy: we prove @{prop"ev m \<Longrightarrow> ev(m - 2)"} by induction on @{prop"ev m"}. |
|
586 |
Case @{thm[source]ev0} requires us to prove @{prop"ev(0 - 2)"}, which follows |
|
587 |
from @{prop"ev 0"} because @{prop"0 - 2 = (0::nat)"} on type @{typ nat}. In |
|
588 |
case @{thm[source]evSS} we have \mbox{@{prop"m = n+(2::nat)"}} and may assume |
|
589 |
@{prop"ev n"}, which implies @{prop"ev (m - 2)"} because @{text"m - 2 = (n + |
|
590 |
2) - 2 = n"}. We did not need the induction hypothesis at all for this proof, |
|
47711 | 591 |
it is just a case analysis of which rule was used, but having @{prop"ev |
47269 | 592 |
n"} at our disposal in case @{thm[source]evSS} was essential. |
47711 | 593 |
This case analysis of rules is also called ``rule inversion'' |
47269 | 594 |
and is discussed in more detail in \autoref{ch:Isar}. |
595 |
||
596 |
\subsubsection{In Isabelle} |
|
597 |
||
598 |
Let us now recast the above informal proofs in Isabelle. For a start, |
|
599 |
we use @{const Suc} terms instead of numerals in rule @{thm[source]evSS}: |
|
600 |
@{thm[display] evSS} |
|
601 |
This avoids the difficulty of unifying @{text"n+2"} with some numeral, |
|
602 |
which is not automatic. |
|
603 |
||
604 |
The simplest way to prove @{prop"ev(Suc(Suc(Suc(Suc 0))))"} is in a forward |
|
605 |
direction: @{text "evSS[OF evSS[OF ev0]]"} yields the theorem @{thm evSS[OF |
|
606 |
evSS[OF ev0]]}. Alternatively, you can also prove it as a lemma in backwards |
|
607 |
fashion. Although this is more verbose, it allows us to demonstrate how each |
|
608 |
rule application changes the proof state: *} |
|
609 |
||
610 |
lemma "ev(Suc(Suc(Suc(Suc 0))))" |
|
611 |
txt{* |
|
612 |
@{subgoals[display,indent=0,goals_limit=1]} |
|
613 |
*} |
|
614 |
apply(rule evSS) |
|
615 |
txt{* |
|
616 |
@{subgoals[display,indent=0,goals_limit=1]} |
|
617 |
*} |
|
618 |
apply(rule evSS) |
|
619 |
txt{* |
|
620 |
@{subgoals[display,indent=0,goals_limit=1]} |
|
621 |
*} |
|
622 |
apply(rule ev0) |
|
623 |
done |
|
624 |
||
625 |
text{* \indent |
|
626 |
Rule induction is applied by giving the induction rule explicitly via the |
|
627 |
@{text"rule:"} modifier:*} |
|
628 |
||
629 |
lemma "ev m \<Longrightarrow> even m" |
|
630 |
apply(induction rule: ev.induct) |
|
631 |
by(simp_all) |
|
632 |
||
633 |
text{* Both cases are automatic. Note that if there are multiple assumptions |
|
634 |
of the form @{prop"ev t"}, method @{text induction} will induct on the leftmost |
|
635 |
one. |
|
636 |
||
637 |
As a bonus, we also prove the remaining direction of the equivalence of |
|
638 |
@{const ev} and @{const even}: |
|
639 |
*} |
|
640 |
||
641 |
lemma "even n \<Longrightarrow> ev n" |
|
642 |
apply(induction n rule: even.induct) |
|
643 |
||
644 |
txt{* This is a proof by computation induction on @{text n} (see |
|
645 |
\autoref{sec:recursive-funs}) that sets up three subgoals corresponding to |
|
646 |
the three equations for @{const even}: |
|
647 |
@{subgoals[display,indent=0]} |
|
648 |
The first and third subgoals follow with @{thm[source]ev0} and @{thm[source]evSS}, and the second subgoal is trivially true because @{prop"even(Suc 0)"} is @{const False}: |
|
649 |
*} |
|
650 |
||
651 |
by (simp_all add: ev0 evSS) |
|
652 |
||
653 |
text{* The rules for @{const ev} make perfect simplification and introduction |
|
654 |
rules because their premises are always smaller than the conclusion. It |
|
655 |
makes sense to turn them into simplification and introduction rules |
|
656 |
permanently, to enhance proof automation: *} |
|
657 |
||
658 |
declare ev.intros[simp,intro] |
|
659 |
||
660 |
text{* The rules of an inductive definition are not simplification rules by |
|
661 |
default because, in contrast to recursive functions, there is no termination |
|
662 |
requirement for inductive definitions. |
|
663 |
||
52361 | 664 |
\subsubsection{Inductive Versus Recursive} |
47269 | 665 |
|
666 |
We have seen two definitions of the notion of evenness, an inductive and a |
|
667 |
recursive one. Which one is better? Much of the time, the recursive one is |
|
668 |
more convenient: it allows us to do rewriting in the middle of terms, and it |
|
669 |
expresses both the positive information (which numbers are even) and the |
|
670 |
negative information (which numbers are not even) directly. An inductive |
|
671 |
definition only expresses the positive information directly. The negative |
|
672 |
information, for example, that @{text 1} is not even, has to be proved from |
|
673 |
it (by induction or rule inversion). On the other hand, rule induction is |
|
47711 | 674 |
tailor-made for proving \mbox{@{prop"ev n \<Longrightarrow> P n"}} because it only asks you |
47269 | 675 |
to prove the positive cases. In the proof of @{prop"even n \<Longrightarrow> P n"} by |
676 |
computation induction via @{thm[source]even.induct}, we are also presented |
|
677 |
with the trivial negative cases. If you want the convenience of both |
|
678 |
rewriting and rule induction, you can make two definitions and show their |
|
679 |
equivalence (as above) or make one definition and prove additional properties |
|
680 |
from it, for example rule induction from computation induction. |
|
681 |
||
682 |
But many concepts do not admit a recursive definition at all because there is |
|
683 |
no datatype for the recursion (for example, the transitive closure of a |
|
47711 | 684 |
relation), or the recursion would not terminate (for example, |
685 |
an interpreter for a programming language). Even if there is a recursive |
|
47269 | 686 |
definition, if we are only interested in the positive information, the |
687 |
inductive definition may be much simpler. |
|
688 |
||
52361 | 689 |
\subsection{The Reflexive Transitive Closure} |
47269 | 690 |
\label{sec:star} |
691 |
||
692 |
Evenness is really more conveniently expressed recursively than inductively. |
|
693 |
As a second and very typical example of an inductive definition we define the |
|
47711 | 694 |
reflexive transitive closure. |
52782 | 695 |
\ifsem |
47711 | 696 |
It will also be an important building block for |
47269 | 697 |
some of the semantics considered in the second part of the book. |
52782 | 698 |
\fi |
47269 | 699 |
|
700 |
The reflexive transitive closure, called @{text star} below, is a function |
|
701 |
that maps a binary predicate to another binary predicate: if @{text r} is of |
|
702 |
type @{text"\<tau> \<Rightarrow> \<tau> \<Rightarrow> bool"} then @{term "star r"} is again of type @{text"\<tau> \<Rightarrow> |
|
703 |
\<tau> \<Rightarrow> bool"}, and @{prop"star r x y"} means that @{text x} and @{text y} are in |
|
704 |
the relation @{term"star r"}. Think @{term"r^*"} when you see @{term"star |
|
705 |
r"}, because @{text"star r"} is meant to be the reflexive transitive closure. |
|
706 |
That is, @{prop"star r x y"} is meant to be true if from @{text x} we can |
|
707 |
reach @{text y} in finitely many @{text r} steps. This concept is naturally |
|
708 |
defined inductively: *} |
|
709 |
||
710 |
inductive star :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" for r where |
|
711 |
refl: "star r x x" | |
|
712 |
step: "r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z" |
|
713 |
||
714 |
text{* The base case @{thm[source] refl} is reflexivity: @{term "x=y"}. The |
|
715 |
step case @{thm[source]step} combines an @{text r} step (from @{text x} to |
|
47711 | 716 |
@{text y}) and a @{term"star r"} step (from @{text y} to @{text z}) into a |
717 |
@{term"star r"} step (from @{text x} to @{text z}). |
|
47269 | 718 |
The ``\isacom{for}~@{text r}'' in the header is merely a hint to Isabelle |
719 |
that @{text r} is a fixed parameter of @{const star}, in contrast to the |
|
720 |
further parameters of @{const star}, which change. As a result, Isabelle |
|
721 |
generates a simpler induction rule. |
|
722 |
||
723 |
By definition @{term"star r"} is reflexive. It is also transitive, but we |
|
724 |
need rule induction to prove that: *} |
|
725 |
||
726 |
lemma star_trans: "star r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z" |
|
727 |
apply(induction rule: star.induct) |
|
728 |
(*<*) |
|
729 |
defer |
|
730 |
apply(rename_tac u x y) |
|
731 |
defer |
|
732 |
(*>*) |
|
54212 | 733 |
txt{* The induction is over @{prop"star r x y"} (the first matching assumption) |
734 |
and we try to prove \mbox{@{prop"star r y z \<Longrightarrow> star r x z"}}, |
|
47269 | 735 |
which we abbreviate by @{prop"P x y"}. These are our two subgoals: |
736 |
@{subgoals[display,indent=0]} |
|
737 |
The first one is @{prop"P x x"}, the result of case @{thm[source]refl}, |
|
738 |
and it is trivial. |
|
739 |
*} |
|
740 |
apply(assumption) |
|
741 |
txt{* Let us examine subgoal @{text 2}, case @{thm[source] step}. |
|
742 |
Assumptions @{prop"r u x"} and \mbox{@{prop"star r x y"}} |
|
743 |
are the premises of rule @{thm[source]step}. |
|
744 |
Assumption @{prop"star r y z \<Longrightarrow> star r x z"} is \mbox{@{prop"P x y"}}, |
|
745 |
the IH coming from @{prop"star r x y"}. We have to prove @{prop"P u y"}, |
|
746 |
which we do by assuming @{prop"star r y z"} and proving @{prop"star r u z"}. |
|
747 |
The proof itself is straightforward: from \mbox{@{prop"star r y z"}} the IH |
|
748 |
leads to @{prop"star r x z"} which, together with @{prop"r u x"}, |
|
749 |
leads to \mbox{@{prop"star r u z"}} via rule @{thm[source]step}: |
|
750 |
*} |
|
751 |
apply(metis step) |
|
752 |
done |
|
753 |
||
754 |
text{* |
|
755 |
||
52361 | 756 |
\subsection{The General Case} |
47269 | 757 |
|
758 |
Inductive definitions have approximately the following general form: |
|
759 |
\begin{quote} |
|
760 |
\isacom{inductive} @{text"I :: \"\<tau> \<Rightarrow> bool\""} \isacom{where} |
|
761 |
\end{quote} |
|
762 |
followed by a sequence of (possibly named) rules of the form |
|
763 |
\begin{quote} |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
764 |
@{text "\<lbrakk> I a\<^sub>1; \<dots>; I a\<^sub>n \<rbrakk> \<Longrightarrow> I a"} |
47269 | 765 |
\end{quote} |
766 |
separated by @{text"|"}. As usual, @{text n} can be 0. |
|
767 |
The corresponding rule induction principle |
|
768 |
@{text I.induct} applies to propositions of the form |
|
769 |
\begin{quote} |
|
770 |
@{prop "I x \<Longrightarrow> P x"} |
|
771 |
\end{quote} |
|
772 |
where @{text P} may itself be a chain of implications. |
|
773 |
\begin{warn} |
|
774 |
Rule induction is always on the leftmost premise of the goal. |
|
775 |
Hence @{text "I x"} must be the first premise. |
|
776 |
\end{warn} |
|
777 |
Proving @{prop "I x \<Longrightarrow> P x"} by rule induction means proving |
|
778 |
for every rule of @{text I} that @{text P} is invariant: |
|
779 |
\begin{quote} |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52782
diff
changeset
|
780 |
@{text "\<lbrakk> I a\<^sub>1; P a\<^sub>1; \<dots>; I a\<^sub>n; P a\<^sub>n \<rbrakk> \<Longrightarrow> P a"} |
47269 | 781 |
\end{quote} |
782 |
||
783 |
The above format for inductive definitions is simplified in a number of |
|
784 |
respects. @{text I} can have any number of arguments and each rule can have |
|
785 |
additional premises not involving @{text I}, so-called \concept{side |
|
786 |
conditions}. In rule inductions, these side-conditions appear as additional |
|
787 |
assumptions. The \isacom{for} clause seen in the definition of the reflexive |
|
788 |
transitive closure merely simplifies the form of the induction rule. |
|
54186 | 789 |
|
790 |
\subsection{Exercises} |
|
791 |
||
54212 | 792 |
\exercise |
793 |
We also could have defined @{const star} as follows: |
|
794 |
*} |
|
795 |
||
796 |
inductive star' :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" for r where |
|
797 |
refl': "star' r x x" | |
|
798 |
step': "star' r x y \<Longrightarrow> r y z \<Longrightarrow> star' r x z" |
|
799 |
||
800 |
text{* |
|
801 |
The single @{text r} step is performer after rather than before the @{text star'} |
|
802 |
steps. Prove @{prop "star' r x y \<Longrightarrow> star r x y"}. You may need a lemma. |
|
803 |
The other direction can also be proved but |
|
804 |
you have to be careful how to formulate the lemma it will require: |
|
805 |
make sure that the assumption about the inductive predicate |
|
806 |
is the first assumption. Assumptions preceding the inductive predicate do not |
|
807 |
take part in the induction. |
|
808 |
\endexercise |
|
809 |
||
810 |
\ifsem |
|
54186 | 811 |
\begin{exercise} |
54203 | 812 |
In \autoref{sec:AExp} we defined a recursive evaluation function |
813 |
@{text "aval :: aexp \<Rightarrow> state \<Rightarrow> val"}. |
|
814 |
Define an inductive evaluation predicate |
|
815 |
@{text "aval_rel :: aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool"} |
|
816 |
and prove that it agrees with the recursive function: |
|
817 |
@{prop "aval_rel a s v \<Longrightarrow> aval a s = v"}, |
|
818 |
@{prop "aval a s = v \<Longrightarrow> aval_rel a s v"} and thus |
|
819 |
\noquotes{@{prop [source] "aval_rel a s v \<longleftrightarrow> aval a s = v"}}. |
|
820 |
\end{exercise} |
|
821 |
||
822 |
\begin{exercise} |
|
54210 | 823 |
Consider the stack machine from Chapter~3 |
824 |
and recall the concept of \concept{stack underflow} |
|
825 |
from Exercise~\ref{exe:stack-underflow}. |
|
826 |
Define an inductive predicate |
|
54186 | 827 |
@{text "ok :: nat \<Rightarrow> instr list \<Rightarrow> nat \<Rightarrow> bool"} |
828 |
such that @{text "ok n is n'"} means that with any initial stack of length |
|
829 |
@{text n} the instructions @{text "is"} can be executed |
|
830 |
without stack underflow and that the final stack has length @{text n'}. |
|
831 |
Prove that @{text ok} correctly computes the final stack size |
|
832 |
@{prop[display] "\<lbrakk>ok n is n'; length stk = n\<rbrakk> \<Longrightarrow> length (exec is s stk) = n'"} |
|
833 |
and that instruction sequences generated by @{text comp} |
|
834 |
cannot cause stack underflow: \ @{text "ok n (comp a) ?"} \ for |
|
835 |
some suitable value of @{text "?"}. |
|
836 |
\end{exercise} |
|
837 |
\fi |
|
47269 | 838 |
*} |
839 |
(*<*) |
|
840 |
end |
|
841 |
(*>*) |