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