| author | wenzelm | 
| Mon, 21 Dec 2020 21:56:20 +0100 | |
| changeset 72974 | 3afd293347cc | 
| parent 69597 | ff784d5a5bfb | 
| child 73511 | 2cdbb6a2f2a7 | 
| permissions | -rw-r--r-- | 
| 47269 | 1  | 
(*<*)  | 
2  | 
theory Isar  | 
|
3  | 
imports LaTeXsugar  | 
|
4  | 
begin  | 
|
| 52059 | 5  | 
declare [[quick_and_dirty]]  | 
| 47269 | 6  | 
(*>*)  | 
| 67406 | 7  | 
text\<open>  | 
| 47269 | 8  | 
Apply-scripts are unreadable and hard to maintain. The language of choice  | 
9  | 
for larger proofs is \concept{Isar}. The two key features of Isar are:
 | 
|
10  | 
\begin{itemize}
 | 
|
11  | 
\item It is structured, not linear.  | 
|
| 56989 | 12  | 
\item It is readable without its being run because  | 
| 47269 | 13  | 
you need to state what you are proving at any given point.  | 
14  | 
\end{itemize}
 | 
|
15  | 
Whereas apply-scripts are like assembly language programs, Isar proofs  | 
|
16  | 
are like structured programs with comments. A typical Isar proof looks like this:  | 
|
| 67406 | 17  | 
\<close>text\<open>  | 
| 47269 | 18  | 
\begin{tabular}{@ {}l}
 | 
19  | 
\isacom{proof}\\
 | 
|
| 69505 | 20  | 
\quad\isacom{assume} \<open>"\<close>$\mathit{formula}_0$\<open>"\<close>\\
 | 
21  | 
\quad\isacom{have} \<open>"\<close>$\mathit{formula}_1$\<open>"\<close> \quad\isacom{by} \<open>simp\<close>\\
 | 
|
| 47269 | 22  | 
\quad\vdots\\  | 
| 69505 | 23  | 
\quad\isacom{have} \<open>"\<close>$\mathit{formula}_n$\<open>"\<close> \quad\isacom{by} \<open>blast\<close>\\
 | 
24  | 
\quad\isacom{show} \<open>"\<close>$\mathit{formula}_{n+1}$\<open>"\<close> \quad\isacom{by} \<open>\<dots>\<close>\\
 | 
|
| 47269 | 25  | 
\isacom{qed}
 | 
26  | 
\end{tabular}
 | 
|
| 67406 | 27  | 
\<close>text\<open>  | 
| 47269 | 28  | 
It proves $\mathit{formula}_0 \Longrightarrow \mathit{formula}_{n+1}$
 | 
| 47704 | 29  | 
(provided each proof step succeeds).  | 
| 47269 | 30  | 
The intermediate \isacom{have} statements are merely stepping stones
 | 
31  | 
on the way towards the \isacom{show} statement that proves the actual
 | 
|
32  | 
goal. In more detail, this is the Isar core syntax:  | 
|
33  | 
\medskip  | 
|
34  | 
||
35  | 
\begin{tabular}{@ {}lcl@ {}}
 | 
|
| 55359 | 36  | 
\textit{proof} &=& \indexed{\isacom{by}}{by} \textit{method}\\
 | 
37  | 
      &$\mid$& \indexed{\isacom{proof}}{proof} [\textit{method}] \ \textit{step}$^*$ \ \indexed{\isacom{qed}}{qed}
 | 
|
| 47269 | 38  | 
\end{tabular}
 | 
39  | 
\medskip  | 
|
40  | 
||
41  | 
\begin{tabular}{@ {}lcl@ {}}
 | 
|
| 55359 | 42  | 
\textit{step} &=& \indexed{\isacom{fix}}{fix} \textit{variables} \\
 | 
43  | 
      &$\mid$& \indexed{\isacom{assume}}{assume} \textit{proposition} \\
 | 
|
44  | 
      &$\mid$& [\indexed{\isacom{from}}{from} \textit{fact}$^+$] (\indexed{\isacom{have}}{have} $\mid$ \indexed{\isacom{show}}{show}) \ \textit{proposition} \ \textit{proof}
 | 
|
| 47269 | 45  | 
\end{tabular}
 | 
46  | 
\medskip  | 
|
47  | 
||
48  | 
\begin{tabular}{@ {}lcl@ {}}
 | 
|
| 69505 | 49  | 
\textit{proposition} &=& [\textit{name}:] \<open>"\<close>\textit{formula}\<open>"\<close>
 | 
| 47269 | 50  | 
\end{tabular}
 | 
51  | 
\medskip  | 
|
52  | 
||
53  | 
\begin{tabular}{@ {}lcl@ {}}
 | 
|
54  | 
\textit{fact} &=& \textit{name} \ $\mid$ \ \dots
 | 
|
55  | 
\end{tabular}
 | 
|
56  | 
\medskip  | 
|
57  | 
||
58  | 
\noindent A proof can either be an atomic \isacom{by} with a single proof
 | 
|
| 69505 | 59  | 
method which must finish off the statement being proved, for example \<open>auto\<close>,  or it can be a \isacom{proof}--\isacom{qed} block of multiple
 | 
| 47269 | 60  | 
steps. Such a block can optionally begin with a proof method that indicates  | 
| 69505 | 61  | 
how to start off the proof, e.g., \mbox{\<open>(induction xs)\<close>}.
 | 
| 47269 | 62  | 
|
63  | 
A step either assumes a proposition or states a proposition  | 
|
64  | 
together with its proof. The optional \isacom{from} clause
 | 
|
65  | 
indicates which facts are to be used in the proof.  | 
|
66  | 
Intermediate propositions are stated with \isacom{have}, the overall goal
 | 
|
| 56989 | 67  | 
is stated with \isacom{show}. A step can also introduce new local variables with
 | 
| 69505 | 68  | 
\isacom{fix}. Logically, \isacom{fix} introduces \<open>\<And>\<close>-quantified
 | 
| 47269 | 69  | 
variables, \isacom{assume} introduces the assumption of an implication
 | 
| 69505 | 70  | 
(\<open>\<Longrightarrow>\<close>) and \isacom{have}/\isacom{show} introduce the conclusion.
 | 
| 47269 | 71  | 
|
72  | 
Propositions are optionally named formulas. These names can be referred to in  | 
|
73  | 
later \isacom{from} clauses. In the simplest case, a fact is such a name.
 | 
|
| 69505 | 74  | 
But facts can also be composed with \<open>OF\<close> and \<open>of\<close> as shown in  | 
| 58602 | 75  | 
\autoref{sec:forward-proof} --- hence the \dots\ in the above grammar.  Note
 | 
| 47269 | 76  | 
that assumptions, intermediate \isacom{have} statements and global lemmas all
 | 
77  | 
have the same status and are thus collectively referred to as  | 
|
| 55317 | 78  | 
\conceptidx{facts}{fact}.
 | 
| 47269 | 79  | 
|
| 69505 | 80  | 
Fact names can stand for whole lists of facts. For example, if \<open>f\<close> is  | 
81  | 
defined by command \isacom{fun}, \<open>f.simps\<close> refers to the whole list of
 | 
|
82  | 
recursion equations defining \<open>f\<close>. Individual facts can be selected by  | 
|
83  | 
writing \<open>f.simps(2)\<close>, whole sublists by writing \<open>f.simps(2-4)\<close>.  | 
|
| 47269 | 84  | 
|
85  | 
||
| 52361 | 86  | 
\section{Isar by Example}
 | 
| 47269 | 87  | 
|
| 47704 | 88  | 
We show a number of proofs of Cantor's theorem that a function from a set to  | 
| 47269 | 89  | 
its powerset cannot be surjective, illustrating various features of Isar. The  | 
| 69597 | 90  | 
constant \<^const>\<open>surj\<close> is predefined.  | 
| 67406 | 91  | 
\<close>  | 
| 47269 | 92  | 
|
93  | 
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"  | 
|
94  | 
proof  | 
|
95  | 
assume 0: "surj f"  | 
|
96  | 
from 0 have 1: "\<forall>A. \<exists>a. A = f a" by(simp add: surj_def)  | 
|
97  | 
  from 1 have 2: "\<exists>a. {x. x \<notin> f x} = f a" by blast
 | 
|
98  | 
from 2 show "False" by blast  | 
|
99  | 
qed  | 
|
100  | 
||
| 67406 | 101  | 
text\<open>  | 
| 56989 | 102  | 
The \isacom{proof} command lacks an explicit method by which to perform
 | 
| 47269 | 103  | 
the proof. In such cases Isabelle tries to use some standard introduction  | 
| 69505 | 104  | 
rule, in the above case for \<open>\<not>\<close>:  | 
| 47269 | 105  | 
\[  | 
106  | 
\inferrule{
 | 
|
107  | 
\mbox{@{thm (prem 1) notI}}}
 | 
|
108  | 
{\mbox{@{thm (concl) notI}}}
 | 
|
109  | 
\]  | 
|
| 69597 | 110  | 
In order to prove \<^prop>\<open>~ P\<close>, assume \<open>P\<close> and show \<open>False\<close>.  | 
| 
64852
 
f3504bc69ea3
fix problems because of "surj" input abbreviation; tuned
 
nipkow 
parents: 
61517 
diff
changeset
 | 
111  | 
Thus we may assume \mbox{\noquotes{@{prop [source] "surj f"}}}. The proof shows that names of propositions
 | 
| 58602 | 112  | 
may be (single!) digits --- meaningful names are hard to invent and are often  | 
| 47269 | 113  | 
not necessary. Both \isacom{have} steps are obvious. The second one introduces
 | 
| 69597 | 114  | 
the diagonal set \<^term>\<open>{x. x \<notin> f x}\<close>, the key idea in the proof.
 | 
| 69505 | 115  | 
If you wonder why \<open>2\<close> directly implies \<open>False\<close>: from \<open>2\<close>  | 
| 69597 | 116  | 
it follows that \<^prop>\<open>a \<notin> f a \<longleftrightarrow> a \<in> f a\<close>.  | 
| 47269 | 117  | 
|
| 69505 | 118  | 
\subsection{\indexed{\<open>this\<close>}{this}, \indexed{\isacom{then}}{then}, \indexed{\isacom{hence}}{hence} and \indexed{\isacom{thus}}{thus}}
 | 
| 47269 | 119  | 
|
120  | 
Labels should be avoided. They interrupt the flow of the reader who has to  | 
|
121  | 
scan the context for the point where the label was introduced. Ideally, the  | 
|
122  | 
proof is a linear flow, where the output of one step becomes the input of the  | 
|
| 58605 | 123  | 
next step, piping the previously proved fact into the next proof, like  | 
| 69505 | 124  | 
in a UNIX pipe. In such cases the predefined name \<open>this\<close> can be used  | 
| 47269 | 125  | 
to refer to the proposition proved in the previous step. This allows us to  | 
126  | 
eliminate all labels from our proof (we suppress the \isacom{lemma} statement):
 | 
|
| 67406 | 127  | 
\<close>  | 
| 47269 | 128  | 
(*<*)  | 
129  | 
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"  | 
|
130  | 
(*>*)  | 
|
131  | 
proof  | 
|
132  | 
assume "surj f"  | 
|
133  | 
  from this have "\<exists>a. {x. x \<notin> f x} = f a" by(auto simp: surj_def)
 | 
|
134  | 
from this show "False" by blast  | 
|
135  | 
qed  | 
|
136  | 
||
| 67406 | 137  | 
text\<open>We have also taken the opportunity to compress the two \isacom{have}
 | 
| 47269 | 138  | 
steps into one.  | 
139  | 
||
140  | 
To compact the text further, Isar has a few convenient abbreviations:  | 
|
141  | 
\medskip  | 
|
142  | 
||
| 54839 | 143  | 
\begin{tabular}{r@ {\quad=\quad}l}
 | 
| 69505 | 144  | 
\isacom{then} & \isacom{from} \<open>this\<close>\\
 | 
| 54839 | 145  | 
\isacom{thus} & \isacom{then} \isacom{show}\\
 | 
146  | 
\isacom{hence} & \isacom{then} \isacom{have}
 | 
|
| 47269 | 147  | 
\end{tabular}
 | 
148  | 
\medskip  | 
|
149  | 
||
150  | 
\noindent  | 
|
151  | 
With the help of these abbreviations the proof becomes  | 
|
| 67406 | 152  | 
\<close>  | 
| 47269 | 153  | 
(*<*)  | 
154  | 
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"  | 
|
155  | 
(*>*)  | 
|
156  | 
proof  | 
|
157  | 
assume "surj f"  | 
|
158  | 
  hence "\<exists>a. {x. x \<notin> f x} = f a" by(auto simp: surj_def)
 | 
|
159  | 
thus "False" by blast  | 
|
160  | 
qed  | 
|
| 67406 | 161  | 
text\<open>  | 
| 47269 | 162  | 
|
163  | 
There are two further linguistic variations:  | 
|
164  | 
\medskip  | 
|
165  | 
||
| 54839 | 166  | 
\begin{tabular}{r@ {\quad=\quad}l}
 | 
| 55359 | 167  | 
(\isacom{have}$\mid$\isacom{show}) \ \textit{prop} \ \indexed{\isacom{using}}{using} \ \textit{facts}
 | 
| 54839 | 168  | 
&  | 
| 47269 | 169  | 
\isacom{from} \ \textit{facts} \ (\isacom{have}$\mid$\isacom{show}) \ \textit{prop}\\
 | 
| 55359 | 170  | 
\indexed{\isacom{with}}{with} \ \textit{facts} & \isacom{from} \ \textit{facts} \isa{this}
 | 
| 47269 | 171  | 
\end{tabular}
 | 
172  | 
\medskip  | 
|
173  | 
||
| 47711 | 174  | 
\noindent The \isacom{using} idiom de-emphasizes the used facts by moving them
 | 
| 47269 | 175  | 
behind the proposition.  | 
176  | 
||
| 55359 | 177  | 
\subsection{Structured Lemma Statements: \indexed{\isacom{fixes}}{fixes}, \indexed{\isacom{assumes}}{assumes}, \indexed{\isacom{shows}}{shows}}
 | 
178  | 
\index{lemma@\isacom{lemma}}
 | 
|
| 47269 | 179  | 
Lemmas can also be stated in a more structured fashion. To demonstrate this  | 
| 
64852
 
f3504bc69ea3
fix problems because of "surj" input abbreviation; tuned
 
nipkow 
parents: 
61517 
diff
changeset
 | 
180  | 
feature with Cantor's theorem, we rephrase \noquotes{@{prop[source]"\<not> surj f"}}
 | 
| 47269 | 181  | 
a little:  | 
| 67406 | 182  | 
\<close>  | 
| 47269 | 183  | 
|
184  | 
lemma  | 
|
185  | 
fixes f :: "'a \<Rightarrow> 'a set"  | 
|
186  | 
assumes s: "surj f"  | 
|
187  | 
shows "False"  | 
|
188  | 
||
| 67406 | 189  | 
txt\<open>The optional \isacom{fixes} part allows you to state the types of
 | 
| 47269 | 190  | 
variables up front rather than by decorating one of their occurrences in the  | 
191  | 
formula with a type constraint. The key advantage of the structured format is  | 
|
| 47306 | 192  | 
the \isacom{assumes} part that allows you to name each assumption; multiple
 | 
193  | 
assumptions can be separated by \isacom{and}. The
 | 
|
| 47269 | 194  | 
\isacom{shows} part gives the goal. The actual theorem that will come out of
 | 
| 
64852
 
f3504bc69ea3
fix problems because of "surj" input abbreviation; tuned
 
nipkow 
parents: 
61517 
diff
changeset
 | 
195  | 
the proof is \noquotes{@{prop[source]"surj f \<Longrightarrow> False"}}, but during the proof the assumption
 | 
| 69505 | 196  | 
\noquotes{@{prop[source]"surj f"}} is available under the name \<open>s\<close> like any other fact.
 | 
| 67406 | 197  | 
\<close>  | 
| 47269 | 198  | 
|
199  | 
proof -  | 
|
200  | 
  have "\<exists> a. {x. x \<notin> f x} = f a" using s
 | 
|
201  | 
by(auto simp: surj_def)  | 
|
202  | 
thus "False" by blast  | 
|
203  | 
qed  | 
|
204  | 
||
| 67406 | 205  | 
text\<open>  | 
| 56312 | 206  | 
\begin{warn}
 | 
207  | 
Note the hyphen after the \isacom{proof} command.
 | 
|
| 56989 | 208  | 
It is the null method that does nothing to the goal. Leaving it out would be asking  | 
| 69597 | 209  | 
Isabelle to try some suitable introduction rule on the goal \<^const>\<open>False\<close> --- but  | 
| 56312 | 210  | 
there is no such rule and \isacom{proof} would fail.
 | 
211  | 
\end{warn}
 | 
|
| 
64852
 
f3504bc69ea3
fix problems because of "surj" input abbreviation; tuned
 
nipkow 
parents: 
61517 
diff
changeset
 | 
212  | 
In the \isacom{have} step the assumption \noquotes{@{prop[source]"surj f"}} is now
 | 
| 69505 | 213  | 
referenced by its name \<open>s\<close>. The duplication of \noquotes{@{prop[source]"surj f"}} in the
 | 
| 47269 | 214  | 
above proofs (once in the statement of the lemma, once in its proof) has been  | 
215  | 
eliminated.  | 
|
216  | 
||
| 47704 | 217  | 
Stating a lemma with \isacom{assumes}-\isacom{shows} implicitly introduces the
 | 
| 69505 | 218  | 
name \indexed{\<open>assms\<close>}{assms} that stands for the list of all assumptions. You can refer
 | 
219  | 
to individual assumptions by \<open>assms(1)\<close>, \<open>assms(2)\<close>, etc.,  | 
|
| 47269 | 220  | 
thus obviating the need to name them individually.  | 
221  | 
||
| 52361 | 222  | 
\section{Proof Patterns}
 | 
| 47269 | 223  | 
|
224  | 
We show a number of important basic proof patterns. Many of them arise from  | 
|
225  | 
the rules of natural deduction that are applied by \isacom{proof} by
 | 
|
226  | 
default. The patterns are phrased in terms of \isacom{show} but work for
 | 
|
227  | 
\isacom{have} and \isacom{lemma}, too.
 | 
|
228  | 
||
| 65348 | 229  | 
\ifsem\else  | 
230  | 
\subsection{Logic}
 | 
|
231  | 
\fi  | 
|
232  | 
||
| 47711 | 233  | 
We start with two forms of \concept{case analysis}:
 | 
| 69505 | 234  | 
starting from a formula \<open>P\<close> we have the two cases \<open>P\<close> and  | 
| 69597 | 235  | 
\<^prop>\<open>~P\<close>, and starting from a fact \<^prop>\<open>P \<or> Q\<close>  | 
| 69505 | 236  | 
we have the two cases \<open>P\<close> and \<open>Q\<close>:  | 
| 67406 | 237  | 
\<close>text_raw\<open>  | 
| 47269 | 238  | 
\begin{tabular}{@ {}ll@ {}}
 | 
239  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
240  | 
\isa{%
 | 
|
| 67406 | 241  | 
\<close>  | 
| 47269 | 242  | 
(*<*)lemma "R" proof-(*>*)  | 
243  | 
show "R"  | 
|
244  | 
proof cases  | 
|
245  | 
assume "P"  | 
|
| 67406 | 246  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
247  | 
show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 248  | 
next  | 
249  | 
assume "\<not> P"  | 
|
| 67406 | 250  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
251  | 
show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 252  | 
qed(*<*)oops(*>*)  | 
| 67406 | 253  | 
text_raw \<open>}  | 
| 69505 | 254  | 
\end{minipage}\index{cases@\<open>cases\<close>}
 | 
| 47269 | 255  | 
&  | 
256  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
257  | 
\isa{%
 | 
|
| 67406 | 258  | 
\<close>  | 
| 47269 | 259  | 
(*<*)lemma "R" proof-(*>*)  | 
| 67406 | 260  | 
have "P \<or> Q" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
| 47269 | 261  | 
then show "R"  | 
262  | 
proof  | 
|
263  | 
assume "P"  | 
|
| 67406 | 264  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
265  | 
show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 266  | 
next  | 
267  | 
assume "Q"  | 
|
| 67406 | 268  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
269  | 
show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 270  | 
qed(*<*)oops(*>*)  | 
271  | 
||
| 67406 | 272  | 
text_raw \<open>}  | 
| 47269 | 273  | 
\end{minipage}
 | 
274  | 
\end{tabular}
 | 
|
275  | 
\medskip  | 
|
276  | 
\begin{isamarkuptext}%
 | 
|
277  | 
How to prove a logical equivalence:  | 
|
278  | 
\end{isamarkuptext}%
 | 
|
279  | 
\isa{%
 | 
|
| 67406 | 280  | 
\<close>  | 
| 47269 | 281  | 
(*<*)lemma "P\<longleftrightarrow>Q" proof-(*>*)  | 
282  | 
show "P \<longleftrightarrow> Q"  | 
|
283  | 
proof  | 
|
284  | 
assume "P"  | 
|
| 67406 | 285  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
286  | 
show "Q" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 287  | 
next  | 
288  | 
assume "Q"  | 
|
| 67406 | 289  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
290  | 
show "P" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 291  | 
qed(*<*)qed(*>*)  | 
| 67406 | 292  | 
text_raw \<open>}  | 
| 47269 | 293  | 
\medskip  | 
294  | 
\begin{isamarkuptext}%
 | 
|
| 55361 | 295  | 
Proofs by contradiction (@{thm[source] ccontr} stands for ``classical contradiction''):
 | 
| 47269 | 296  | 
\end{isamarkuptext}%
 | 
297  | 
\begin{tabular}{@ {}ll@ {}}
 | 
|
298  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
299  | 
\isa{%
 | 
|
| 67406 | 300  | 
\<close>  | 
| 47269 | 301  | 
(*<*)lemma "\<not> P" proof-(*>*)  | 
302  | 
show "\<not> P"  | 
|
303  | 
proof  | 
|
304  | 
assume "P"  | 
|
| 67406 | 305  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
306  | 
show "False" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 307  | 
qed(*<*)oops(*>*)  | 
308  | 
||
| 67406 | 309  | 
text_raw \<open>}  | 
| 47269 | 310  | 
\end{minipage}
 | 
311  | 
&  | 
|
312  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
313  | 
\isa{%
 | 
|
| 67406 | 314  | 
\<close>  | 
| 47269 | 315  | 
(*<*)lemma "P" proof-(*>*)  | 
316  | 
show "P"  | 
|
317  | 
proof (rule ccontr)  | 
|
318  | 
assume "\<not>P"  | 
|
| 67406 | 319  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
320  | 
show "False" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 321  | 
qed(*<*)oops(*>*)  | 
322  | 
||
| 67406 | 323  | 
text_raw \<open>}  | 
| 47269 | 324  | 
\end{minipage}
 | 
325  | 
\end{tabular}
 | 
|
326  | 
\medskip  | 
|
327  | 
\begin{isamarkuptext}%
 | 
|
328  | 
How to prove quantified formulas:  | 
|
329  | 
\end{isamarkuptext}%
 | 
|
330  | 
\begin{tabular}{@ {}ll@ {}}
 | 
|
331  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
332  | 
\isa{%
 | 
|
| 67406 | 333  | 
\<close>  | 
| 67613 | 334  | 
(*<*)lemma "\<forall>x. P x" proof-(*>*)  | 
| 47269 | 335  | 
show "\<forall>x. P(x)"  | 
336  | 
proof  | 
|
337  | 
fix x  | 
|
| 67406 | 338  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
339  | 
show "P(x)" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 340  | 
qed(*<*)oops(*>*)  | 
341  | 
||
| 67406 | 342  | 
text_raw \<open>}  | 
| 47269 | 343  | 
\end{minipage}
 | 
344  | 
&  | 
|
345  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
346  | 
\isa{%
 | 
|
| 67406 | 347  | 
\<close>  | 
| 67613 | 348  | 
(*<*)lemma "\<exists>x. P(x)" proof-(*>*)  | 
| 47269 | 349  | 
show "\<exists>x. P(x)"  | 
350  | 
proof  | 
|
| 67406 | 351  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
352  | 
show "P(witness)" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 353  | 
qed  | 
354  | 
(*<*)oops(*>*)  | 
|
355  | 
||
| 67406 | 356  | 
text_raw \<open>}  | 
| 47269 | 357  | 
\end{minipage}
 | 
358  | 
\end{tabular}
 | 
|
359  | 
\medskip  | 
|
360  | 
\begin{isamarkuptext}%
 | 
|
361  | 
In the proof of \noquotes{@{prop[source]"\<forall>x. P(x)"}},
 | 
|
| 69505 | 362  | 
the step \indexed{\isacom{fix}}{fix}~\<open>x\<close> introduces a locally fixed variable \<open>x\<close>
 | 
| 47269 | 363  | 
into the subproof, the proverbial ``arbitrary but fixed value''.  | 
| 69505 | 364  | 
Instead of \<open>x\<close> we could have chosen any name in the subproof.  | 
| 47269 | 365  | 
In the proof of \noquotes{@{prop[source]"\<exists>x. P(x)"}},
 | 
| 69505 | 366  | 
\<open>witness\<close> is some arbitrary  | 
367  | 
term for which we can prove that it satisfies \<open>P\<close>.  | 
|
| 47269 | 368  | 
|
369  | 
How to reason forward from \noquotes{@{prop[source] "\<exists>x. P(x)"}}:
 | 
|
370  | 
\end{isamarkuptext}%
 | 
|
| 67406 | 371  | 
\<close>  | 
| 67613 | 372  | 
(*<*)lemma True proof- assume 1: "\<exists>x. P x"(*>*)  | 
| 67406 | 373  | 
have "\<exists>x. P(x)" (*<*)by(rule 1)(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
| 47269 | 374  | 
then obtain x where p: "P(x)" by blast  | 
375  | 
(*<*)oops(*>*)  | 
|
| 67406 | 376  | 
text\<open>  | 
| 69505 | 377  | 
After the \indexed{\isacom{obtain}}{obtain} step, \<open>x\<close> (we could have chosen any name)
 | 
| 47269 | 378  | 
is a fixed local  | 
| 69505 | 379  | 
variable, and \<open>p\<close> is the name of the fact  | 
| 47269 | 380  | 
\noquotes{@{prop[source] "P(x)"}}.
 | 
| 69505 | 381  | 
This pattern works for one or more \<open>x\<close>.  | 
| 47269 | 382  | 
As an example of the \isacom{obtain} command, here is the proof of
 | 
383  | 
Cantor's theorem in more detail:  | 
|
| 67406 | 384  | 
\<close>  | 
| 47269 | 385  | 
|
386  | 
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"  | 
|
387  | 
proof  | 
|
388  | 
assume "surj f"  | 
|
389  | 
  hence  "\<exists>a. {x. x \<notin> f x} = f a" by(auto simp: surj_def)
 | 
|
390  | 
  then obtain a where  "{x. x \<notin> f x} = f a"  by blast
 | 
|
391  | 
hence "a \<notin> f a \<longleftrightarrow> a \<in> f a" by blast  | 
|
392  | 
thus "False" by blast  | 
|
393  | 
qed  | 
|
394  | 
||
| 67406 | 395  | 
text_raw\<open>  | 
| 47269 | 396  | 
\begin{isamarkuptext}%
 | 
| 47306 | 397  | 
|
398  | 
Finally, how to prove set equality and subset relationship:  | 
|
| 47269 | 399  | 
\end{isamarkuptext}%
 | 
400  | 
\begin{tabular}{@ {}ll@ {}}
 | 
|
401  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
402  | 
\isa{%
 | 
|
| 67406 | 403  | 
\<close>  | 
| 47269 | 404  | 
(*<*)lemma "A = (B::'a set)" proof-(*>*)  | 
405  | 
show "A = B"  | 
|
406  | 
proof  | 
|
| 67406 | 407  | 
show "A \<subseteq> B" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
| 47269 | 408  | 
next  | 
| 67406 | 409  | 
show "B \<subseteq> A" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
| 47269 | 410  | 
qed(*<*)qed(*>*)  | 
411  | 
||
| 67406 | 412  | 
text_raw \<open>}  | 
| 47269 | 413  | 
\end{minipage}
 | 
414  | 
&  | 
|
415  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
416  | 
\isa{%
 | 
|
| 67406 | 417  | 
\<close>  | 
| 47269 | 418  | 
(*<*)lemma "A <= (B::'a set)" proof-(*>*)  | 
419  | 
show "A \<subseteq> B"  | 
|
420  | 
proof  | 
|
421  | 
fix x  | 
|
422  | 
assume "x \<in> A"  | 
|
| 67406 | 423  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
424  | 
show "x \<in> B" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 425  | 
qed(*<*)qed(*>*)  | 
426  | 
||
| 67406 | 427  | 
text_raw \<open>}  | 
| 47269 | 428  | 
\end{minipage}
 | 
429  | 
\end{tabular}
 | 
|
430  | 
\begin{isamarkuptext}%
 | 
|
| 65348 | 431  | 
|
432  | 
\ifsem\else  | 
|
433  | 
\subsection{Chains of (In)Equations}
 | 
|
434  | 
||
435  | 
In textbooks, chains of equations (and inequations) are often displayed like this:  | 
|
436  | 
\begin{quote}
 | 
|
437  | 
\begin{tabular}{@ {}l@ {\qquad}l@ {}}
 | 
|
438  | 
$t_1 = t_2$ & \isamath{\,\langle\mathit{justification}\rangle}\\
 | 
|
439  | 
$\phantom{t_1} = t_3$ & \isamath{\,\langle\mathit{justification}\rangle}\\
 | 
|
440  | 
\quad $\vdots$\\  | 
|
441  | 
$\phantom{t_1} = t_n$ & \isamath{\,\langle\mathit{justification}\rangle}
 | 
|
442  | 
\end{tabular}
 | 
|
443  | 
\end{quote}
 | 
|
444  | 
The Isar equivalent is this:  | 
|
445  | 
||
446  | 
\begin{samepage}
 | 
|
447  | 
\begin{quote}
 | 
|
448  | 
\isakeyword{have} \<open>"t\<^sub>1 = t\<^sub>2"\<close> \isasymproof\\
 | 
|
449  | 
\isakeyword{also have} \<open>"... = t\<^sub>3"\<close> \isasymproof\\
 | 
|
450  | 
\quad $\vdots$\\  | 
|
451  | 
\isakeyword{also have} \<open>"... = t\<^sub>n"\<close> \isasymproof \\
 | 
|
| 65349 | 452  | 
\isakeyword{finally show} \<open>"t\<^sub>1 = t\<^sub>n"\<close>\ \texttt{.}
 | 
| 65348 | 453  | 
\end{quote}
 | 
454  | 
\end{samepage}
 | 
|
455  | 
||
456  | 
\noindent  | 
|
457  | 
The ``\<open>...\<close>'' and ``\<open>.\<close>'' deserve some explanation:  | 
|
458  | 
\begin{description}
 | 
|
459  | 
\item[``\<open>...\<close>''] is literally three dots. It is the name of an unknown that Isar  | 
|
460  | 
automatically instantiates with the right-hand side of the previous equation.  | 
|
| 69597 | 461  | 
In general, if \<open>this\<close> is the theorem \<^term>\<open>p t\<^sub>1 t\<^sub>2\<close> then ``\<open>...\<close>''  | 
| 65348 | 462  | 
stands for \<open>t\<^sub>2\<close>.  | 
| 65349 | 463  | 
\item[``\<open>.\<close>''] (a single dot) is a proof method that solves a goal by one of the  | 
| 65352 | 464  | 
assumptions. This works here because the result of \isakeyword{finally}
 | 
| 65349 | 465  | 
is the theorem \mbox{\<open>t\<^sub>1 = t\<^sub>n\<close>},
 | 
466  | 
\isakeyword{show} \<open>"t\<^sub>1 = t\<^sub>n"\<close> states the theorem explicitly,
 | 
|
467  | 
and ``\<open>.\<close>'' proves the theorem with the result of \isakeyword{finally}.
 | 
|
| 65348 | 468  | 
\end{description}
 | 
469  | 
The above proof template also works for arbitrary mixtures of \<open>=\<close>, \<open>\<le>\<close> and \<open><\<close>,  | 
|
470  | 
for example:  | 
|
471  | 
\begin{quote}
 | 
|
472  | 
\isakeyword{have} \<open>"t\<^sub>1 < t\<^sub>2"\<close> \isasymproof\\
 | 
|
473  | 
\isakeyword{also have} \<open>"... = t\<^sub>3"\<close> \isasymproof\\
 | 
|
474  | 
\quad $\vdots$\\  | 
|
475  | 
\isakeyword{also have} \<open>"... \<le> t\<^sub>n"\<close> \isasymproof \\
 | 
|
| 65349 | 476  | 
\isakeyword{finally show} \<open>"t\<^sub>1 < t\<^sub>n"\<close>\ \texttt{.}
 | 
| 65348 | 477  | 
\end{quote}
 | 
| 65349 | 478  | 
The relation symbol in the \isakeyword{finally} step needs to be the most precise one
 | 
479  | 
possible. In the example above, you must not write \<open>t\<^sub>1 \<le> t\<^sub>n\<close> instead of \mbox{\<open>t\<^sub>1 < t\<^sub>n\<close>}.
 | 
|
| 65348 | 480  | 
|
481  | 
\begin{warn}
 | 
|
482  | 
Isabelle only supports \<open>=\<close>, \<open>\<le>\<close> and \<open><\<close> but not \<open>\<ge>\<close> and \<open>>\<close>  | 
|
483  | 
in (in)equation chains (by default).  | 
|
484  | 
\end{warn}
 | 
|
485  | 
||
486  | 
If you want to go beyond merely using the above proof patterns and want to  | 
|
487  | 
understand what \isakeyword{also} and \isakeyword{finally} mean, read on.
 | 
|
488  | 
There is an Isar theorem variable called \<open>calculation\<close>, similar to \<open>this\<close>.  | 
|
489  | 
When the first \isakeyword{also} in a chain is encountered, Isabelle sets
 | 
|
490  | 
\<open>calculation := this\<close>. In each subsequent \isakeyword{also} step,
 | 
|
491  | 
Isabelle composes the theorems \<open>calculation\<close> and \<open>this\<close> (i.e.\ the two previous  | 
|
492  | 
(in)equalities) using some predefined set of rules including transitivity  | 
|
| 69597 | 493  | 
of \<open>=\<close>, \<open>\<le>\<close> and \<open><\<close> but also mixed rules like \<^prop>\<open>\<lbrakk> x \<le> y; y < z \<rbrakk> \<Longrightarrow> x < z\<close>.  | 
| 65348 | 494  | 
The result of this composition is assigned to \<open>calculation\<close>. Consider  | 
495  | 
\begin{quote}
 | 
|
496  | 
\isakeyword{have} \<open>"t\<^sub>1 \<le> t\<^sub>2"\<close> \isasymproof\\
 | 
|
497  | 
\isakeyword{also} \isakeyword{have} \<open>"... < t\<^sub>3"\<close> \isasymproof\\
 | 
|
498  | 
\isakeyword{also} \isakeyword{have} \<open>"... = t\<^sub>4"\<close> \isasymproof\\
 | 
|
| 65349 | 499  | 
\isakeyword{finally show} \<open>"t\<^sub>1 < t\<^sub>4"\<close>\ \texttt{.}
 | 
| 65348 | 500  | 
\end{quote}
 | 
501  | 
After the first \isakeyword{also}, \<open>calculation\<close> is \<open>"t\<^sub>1 \<le> t\<^sub>2"\<close>,
 | 
|
502  | 
and after the second \isakeyword{also}, \<open>calculation\<close> is \<open>"t\<^sub>1 < t\<^sub>3"\<close>.
 | 
|
503  | 
The command \isakeyword{finally} is short for \isakeyword{also from} \<open>calculation\<close>.
 | 
|
504  | 
Therefore the \isakeyword{also} hidden in \isakeyword{finally} sets \<open>calculation\<close>
 | 
|
505  | 
to \<open>t\<^sub>1 < t\<^sub>4\<close> and the final ``\texttt{.}'' succeeds.
 | 
|
506  | 
||
| 67299 | 507  | 
For more information on this style of proof see @{cite "BauerW-TPHOLs01"}.
 | 
| 65348 | 508  | 
\fi  | 
509  | 
||
| 52361 | 510  | 
\section{Streamlining Proofs}
 | 
| 47269 | 511  | 
|
| 52361 | 512  | 
\subsection{Pattern Matching and Quotations}
 | 
| 47269 | 513  | 
|
514  | 
In the proof patterns shown above, formulas are often duplicated.  | 
|
515  | 
This can make the text harder to read, write and maintain. Pattern matching  | 
|
516  | 
is an abbreviation mechanism to avoid such duplication. Writing  | 
|
517  | 
\begin{quote}
 | 
|
| 69505 | 518  | 
\isacom{show} \ \textit{formula} \<open>(\<close>\indexed{\isacom{is}}{is} \textit{pattern}\<open>)\<close>
 | 
| 47269 | 519  | 
\end{quote}
 | 
520  | 
matches the pattern against the formula, thus instantiating the unknowns in  | 
|
521  | 
the pattern for later use. As an example, consider the proof pattern for  | 
|
| 69505 | 522  | 
\<open>\<longleftrightarrow>\<close>:  | 
| 47269 | 523  | 
\end{isamarkuptext}%
 | 
| 67406 | 524  | 
\<close>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52718 
diff
changeset
 | 
525  | 
(*<*)lemma "formula\<^sub>1 \<longleftrightarrow> formula\<^sub>2" proof-(*>*)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52718 
diff
changeset
 | 
526  | 
show "formula\<^sub>1 \<longleftrightarrow> formula\<^sub>2" (is "?L \<longleftrightarrow> ?R")  | 
| 47269 | 527  | 
proof  | 
528  | 
assume "?L"  | 
|
| 67406 | 529  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
530  | 
show "?R" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 531  | 
next  | 
532  | 
assume "?R"  | 
|
| 67406 | 533  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
534  | 
show "?L" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 535  | 
qed(*<*)qed(*>*)  | 
536  | 
||
| 69505 | 537  | 
text\<open>Instead of duplicating \<open>formula\<^sub>i\<close> in the text, we introduce  | 
538  | 
the two abbreviations \<open>?L\<close> and \<open>?R\<close> by pattern matching.  | 
|
| 47269 | 539  | 
Pattern matching works wherever a formula is stated, in particular  | 
540  | 
with \isacom{have} and \isacom{lemma}.
 | 
|
541  | 
||
| 69505 | 542  | 
The unknown \indexed{\<open>?thesis\<close>}{thesis} is implicitly matched against any goal stated by
 | 
| 67406 | 543  | 
\isacom{lemma} or \isacom{show}. Here is a typical example:\<close>
 | 
| 47269 | 544  | 
|
545  | 
lemma "formula"  | 
|
546  | 
proof -  | 
|
| 67406 | 547  | 
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
 | 
548  | 
show ?thesis (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 549  | 
qed  | 
550  | 
||
| 67406 | 551  | 
text\<open>  | 
| 55359 | 552  | 
Unknowns can also be instantiated with \indexed{\isacom{let}}{let} commands
 | 
| 47269 | 553  | 
\begin{quote}
 | 
| 69505 | 554  | 
\isacom{let} \<open>?t\<close> = \<open>"\<close>\textit{some-big-term}\<open>"\<close>
 | 
| 47269 | 555  | 
\end{quote}
 | 
| 69505 | 556  | 
Later proof steps can refer to \<open>?t\<close>:  | 
| 47269 | 557  | 
\begin{quote}
 | 
| 69505 | 558  | 
\isacom{have} \<open>"\<close>\dots \<open>?t\<close> \dots\<open>"\<close>
 | 
| 47269 | 559  | 
\end{quote}
 | 
560  | 
\begin{warn}
 | 
|
| 69505 | 561  | 
Names of facts are introduced with \<open>name:\<close> and refer to proved  | 
562  | 
theorems. Unknowns \<open>?X\<close> refer to terms or formulas.  | 
|
| 47269 | 563  | 
\end{warn}
 | 
564  | 
||
565  | 
Although abbreviations shorten the text, the reader needs to remember what  | 
|
| 69505 | 566  | 
they stand for. Similarly for names of facts. Names like \<open>1\<close>, \<open>2\<close>  | 
567  | 
and \<open>3\<close> are not helpful and should only be used in short proofs. For  | 
|
| 47704 | 568  | 
longer proofs, descriptive names are better. But look at this example:  | 
| 47269 | 569  | 
\begin{quote}
 | 
| 69505 | 570  | 
\isacom{have} \ \<open>x_gr_0: "x > 0"\<close>\\
 | 
| 47269 | 571  | 
$\vdots$\\  | 
| 69505 | 572  | 
\isacom{from} \<open>x_gr_0\<close> \dots
 | 
| 47269 | 573  | 
\end{quote}
 | 
| 56989 | 574  | 
The name is longer than the fact it stands for! Short facts do not need names;  | 
| 47269 | 575  | 
one can refer to them easily by quoting them:  | 
576  | 
\begin{quote}
 | 
|
| 69505 | 577  | 
\isacom{have} \ \<open>"x > 0"\<close>\\
 | 
| 47269 | 578  | 
$\vdots$\\  | 
| 69505 | 579  | 
\isacom{from} \<open>`x>0`\<close> \dots\index{$IMP053@\<open>`...`\<close>}
 | 
| 47269 | 580  | 
\end{quote}
 | 
| 69505 | 581  | 
Note that the quotes around \<open>x>0\<close> are \conceptnoidx{back quotes}.
 | 
| 47269 | 582  | 
They refer to the fact not by name but by value.  | 
583  | 
||
| 55359 | 584  | 
\subsection{\indexed{\isacom{moreover}}{moreover}}
 | 
585  | 
\index{ultimately@\isacom{ultimately}}
 | 
|
| 47269 | 586  | 
|
587  | 
Sometimes one needs a number of facts to enable some deduction. Of course  | 
|
588  | 
one can name these facts individually, as shown on the right,  | 
|
589  | 
but one can also combine them with \isacom{moreover}, as shown on the left:
 | 
|
| 67406 | 590  | 
\<close>text_raw\<open>  | 
| 47269 | 591  | 
\begin{tabular}{@ {}ll@ {}}
 | 
592  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
593  | 
\isa{%
 | 
|
| 67406 | 594  | 
\<close>  | 
| 47269 | 595  | 
(*<*)lemma "P" proof-(*>*)  | 
| 67406 | 596  | 
have "P\<^sub>1" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
597  | 
moreover have "P\<^sub>2" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 598  | 
moreover  | 
| 67406 | 599  | 
text_raw\<open>\\$\vdots$\\\hspace{-1.4ex}\<close>(*<*)have "True" ..(*>*)
 | 
600  | 
moreover have "P\<^sub>n" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
601  | 
ultimately have "P" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 602  | 
(*<*)oops(*>*)  | 
603  | 
||
| 67406 | 604  | 
text_raw \<open>}  | 
| 47269 | 605  | 
\end{minipage}
 | 
606  | 
&  | 
|
607  | 
\qquad  | 
|
608  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
609  | 
\isa{%
 | 
|
| 67406 | 610  | 
\<close>  | 
| 47269 | 611  | 
(*<*)lemma "P" proof-(*>*)  | 
| 67406 | 612  | 
have lab\<^sub>1: "P\<^sub>1" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
613  | 
have lab\<^sub>2: "P\<^sub>2" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\<close>  | 
|
614  | 
text_raw\<open>\\$\vdots$\\\hspace{-1.4ex}\<close>
 | 
|
615  | 
have lab\<^sub>n: "P\<^sub>n" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
616  | 
from lab\<^sub>1 lab\<^sub>2 text_raw\<open>\ $\dots$\\\<close>  | 
|
617  | 
have "P" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 618  | 
(*<*)oops(*>*)  | 
619  | 
||
| 67406 | 620  | 
text_raw \<open>}  | 
| 47269 | 621  | 
\end{minipage}
 | 
622  | 
\end{tabular}
 | 
|
623  | 
\begin{isamarkuptext}%
 | 
|
| 65348 | 624  | 
The \isacom{moreover} version is no shorter but expresses the structure
 | 
625  | 
a bit more clearly and avoids new names.  | 
|
| 47269 | 626  | 
|
| 67039 | 627  | 
\subsection{Local Lemmas}
 | 
| 47269 | 628  | 
|
| 56989 | 629  | 
Sometimes one would like to prove some lemma locally within a proof,  | 
630  | 
a lemma that shares the current context of assumptions but that  | 
|
| 58502 | 631  | 
has its own assumptions and is generalized over its locally fixed  | 
| 67039 | 632  | 
variables at the end. This is simply an extension of the basic  | 
633  | 
\indexed{\isacom{have}}{have} construct:
 | 
|
634  | 
\begin{quote}
 | 
|
| 69505 | 635  | 
\indexed{\isacom{have}}{have} \<open>B\<close>\
 | 
636  | 
 \indexed{\isacom{if}}{if} \<open>name:\<close> \<open>A\<^sub>1 \<dots> A\<^sub>m\<close>\
 | 
|
637  | 
 \indexed{\isacom{for}}{for} \<open>x\<^sub>1 \<dots> x\<^sub>n\<close>\\
 | 
|
| 67039 | 638  | 
\isasymproof  | 
| 47269 | 639  | 
\end{quote}
 | 
| 69505 | 640  | 
proves \<open>\<lbrakk> A\<^sub>1; \<dots> ; A\<^sub>m \<rbrakk> \<Longrightarrow> B\<close>  | 
641  | 
where all \<open>x\<^sub>i\<close> have been replaced by unknowns \<open>?x\<^sub>i\<close>.  | 
|
| 47269 | 642  | 
As an example we prove a simple fact about divisibility on integers.  | 
| 69505 | 643  | 
The definition of \<open>dvd\<close> is @{thm dvd_def}.
 | 
| 47269 | 644  | 
\end{isamarkuptext}%
 | 
| 67406 | 645  | 
\<close>  | 
| 47269 | 646  | 
|
647  | 
lemma fixes a b :: int assumes "b dvd (a+b)" shows "b dvd a"  | 
|
648  | 
proof -  | 
|
| 67039 | 649  | 
have "\<exists>k'. a = b*k'" if asm: "a+b = b*k" for k  | 
650  | 
proof  | 
|
651  | 
show "a = b*(k - 1)" using asm by(simp add: algebra_simps)  | 
|
652  | 
qed  | 
|
| 47269 | 653  | 
then show ?thesis using assms by(auto simp add: dvd_def)  | 
654  | 
qed  | 
|
655  | 
||
| 67406 | 656  | 
text\<open>  | 
| 47269 | 657  | 
|
| 54436 | 658  | 
\subsection*{Exercises}
 | 
| 52706 | 659  | 
|
| 52661 | 660  | 
\exercise  | 
661  | 
Give a readable, structured proof of the following lemma:  | 
|
| 67406 | 662  | 
\<close>  | 
| 54218 | 663  | 
lemma assumes T: "\<forall>x y. T x y \<or> T y x"  | 
664  | 
and A: "\<forall>x y. A x y \<and> A y x \<longrightarrow> x = y"  | 
|
665  | 
and TA: "\<forall>x y. T x y \<longrightarrow> A x y" and "A x y"  | 
|
666  | 
shows "T x y"  | 
|
| 52661 | 667  | 
(*<*)oops(*>*)  | 
| 67406 | 668  | 
text\<open>  | 
| 52661 | 669  | 
\endexercise  | 
670  | 
||
| 52706 | 671  | 
\exercise  | 
672  | 
Give a readable, structured proof of the following lemma:  | 
|
| 67406 | 673  | 
\<close>  | 
| 67039 | 674  | 
lemma "\<exists>ys zs. xs = ys @ zs \<and>  | 
675  | 
(length ys = length zs \<or> length ys = length zs + 1)"  | 
|
| 52706 | 676  | 
(*<*)oops(*>*)  | 
| 67406 | 677  | 
text\<open>  | 
| 52706 | 678  | 
Hint: There are predefined functions @{const_typ take} and @{const_typ drop}
 | 
| 69505 | 679  | 
such that \<open>take k [x\<^sub>1,\<dots>] = [x\<^sub>1,\<dots>,x\<^sub>k]\<close> and  | 
680  | 
\<open>drop k [x\<^sub>1,\<dots>] = [x\<^bsub>k+1\<^esub>,\<dots>]\<close>. Let sledgehammer find and apply  | 
|
| 69597 | 681  | 
the relevant \<^const>\<open>take\<close> and \<^const>\<open>drop\<close> lemmas for you.  | 
| 52706 | 682  | 
\endexercise  | 
683  | 
||
| 54218 | 684  | 
|
| 52361 | 685  | 
\section{Case Analysis and Induction}
 | 
| 55361 | 686  | 
|
| 52361 | 687  | 
\subsection{Datatype Case Analysis}
 | 
| 55361 | 688  | 
\index{case analysis|(}
 | 
| 47269 | 689  | 
|
| 47711 | 690  | 
We have seen case analysis on formulas. Now we want to distinguish  | 
| 69597 | 691  | 
which form some term takes: is it \<open>0\<close> or of the form \<^term>\<open>Suc n\<close>,  | 
692  | 
is it \<^term>\<open>[]\<close> or of the form \<^term>\<open>x#xs\<close>, etc. Here is a typical example  | 
|
| 69505 | 693  | 
proof by case analysis on the form of \<open>xs\<close>:  | 
| 67406 | 694  | 
\<close>  | 
| 47269 | 695  | 
|
696  | 
lemma "length(tl xs) = length xs - 1"  | 
|
697  | 
proof (cases xs)  | 
|
698  | 
assume "xs = []"  | 
|
699  | 
thus ?thesis by simp  | 
|
700  | 
next  | 
|
701  | 
fix y ys assume "xs = y#ys"  | 
|
702  | 
thus ?thesis by simp  | 
|
703  | 
qed  | 
|
704  | 
||
| 69505 | 705  | 
text\<open>\index{cases@\<open>cases\<close>|(}Function \<open>tl\<close> (''tail'') is defined by @{thm list.sel(2)} and
 | 
| 69597 | 706  | 
@{thm list.sel(3)}. Note that the result type of \<^const>\<open>length\<close> is \<^typ>\<open>nat\<close>
 | 
707  | 
and \<^prop>\<open>0 - 1 = (0::nat)\<close>.  | 
|
| 47269 | 708  | 
|
| 69505 | 709  | 
This proof pattern works for any term \<open>t\<close> whose type is a datatype.  | 
710  | 
The goal has to be proved for each constructor \<open>C\<close>:  | 
|
| 47269 | 711  | 
\begin{quote}
 | 
| 69505 | 712  | 
\isacom{fix} \ \<open>x\<^sub>1 \<dots> x\<^sub>n\<close> \isacom{assume} \<open>"t = C x\<^sub>1 \<dots> x\<^sub>n"\<close>
 | 
| 55361 | 713  | 
\end{quote}\index{case@\isacom{case}|(}
 | 
| 47269 | 714  | 
Each case can be written in a more compact form by means of the \isacom{case}
 | 
715  | 
command:  | 
|
716  | 
\begin{quote}
 | 
|
| 69505 | 717  | 
\isacom{case} \<open>(C x\<^sub>1 \<dots> x\<^sub>n)\<close>
 | 
| 47269 | 718  | 
\end{quote}
 | 
| 47704 | 719  | 
This is equivalent to the explicit \isacom{fix}-\isacom{assume} line
 | 
| 69505 | 720  | 
but also gives the assumption \<open>"t = C x\<^sub>1 \<dots> x\<^sub>n"\<close> a name: \<open>C\<close>,  | 
| 47269 | 721  | 
like the constructor.  | 
722  | 
Here is the \isacom{case} version of the proof above:
 | 
|
| 67406 | 723  | 
\<close>  | 
| 47269 | 724  | 
(*<*)lemma "length(tl xs) = length xs - 1"(*>*)  | 
725  | 
proof (cases xs)  | 
|
726  | 
case Nil  | 
|
727  | 
thus ?thesis by simp  | 
|
728  | 
next  | 
|
729  | 
case (Cons y ys)  | 
|
730  | 
thus ?thesis by simp  | 
|
731  | 
qed  | 
|
732  | 
||
| 69505 | 733  | 
text\<open>Remember that \<open>Nil\<close> and \<open>Cons\<close> are the alphanumeric names  | 
734  | 
for \<open>[]\<close> and \<open>#\<close>. The names of the assumptions  | 
|
| 47269 | 735  | 
are not used because they are directly piped (via \isacom{thus})
 | 
736  | 
into the proof of the claim.  | 
|
| 55361 | 737  | 
\index{case analysis|)}
 | 
| 47269 | 738  | 
|
| 52361 | 739  | 
\subsection{Structural Induction}
 | 
| 55361 | 740  | 
\index{induction|(}
 | 
741  | 
\index{structural induction|(}
 | 
|
| 47269 | 742  | 
|
743  | 
We illustrate structural induction with an example based on natural numbers:  | 
|
| 69505 | 744  | 
the sum (\<open>\<Sum>\<close>) of the first \<open>n\<close> natural numbers  | 
| 69597 | 745  | 
(\<open>{0..n::nat}\<close>) is equal to \mbox{\<^term>\<open>n*(n+1) div 2::nat\<close>}.
 | 
| 47269 | 746  | 
Never mind the details, just focus on the pattern:  | 
| 67406 | 747  | 
\<close>  | 
| 47269 | 748  | 
|
| 47711 | 749  | 
lemma "\<Sum>{0..n::nat} = n*(n+1) div 2"
 | 
| 47269 | 750  | 
proof (induction n)  | 
751  | 
  show "\<Sum>{0..0::nat} = 0*(0+1) div 2" by simp
 | 
|
752  | 
next  | 
|
753  | 
  fix n assume "\<Sum>{0..n::nat} = n*(n+1) div 2"
 | 
|
| 47711 | 754  | 
  thus "\<Sum>{0..Suc n} = Suc n*(Suc n+1) div 2" by simp
 | 
| 47269 | 755  | 
qed  | 
756  | 
||
| 67406 | 757  | 
text\<open>Except for the rewrite steps, everything is explicitly given. This  | 
| 47269 | 758  | 
makes the proof easily readable, but the duplication means it is tedious to  | 
759  | 
write and maintain. Here is how pattern  | 
|
| 67406 | 760  | 
matching can completely avoid any duplication:\<close>  | 
| 47269 | 761  | 
|
762  | 
lemma "\<Sum>{0..n::nat} = n*(n+1) div 2" (is "?P n")
 | 
|
763  | 
proof (induction n)  | 
|
764  | 
show "?P 0" by simp  | 
|
765  | 
next  | 
|
766  | 
fix n assume "?P n"  | 
|
767  | 
thus "?P(Suc n)" by simp  | 
|
768  | 
qed  | 
|
769  | 
||
| 69505 | 770  | 
text\<open>The first line introduces an abbreviation \<open>?P n\<close> for the goal.  | 
771  | 
Pattern matching \<open>?P n\<close> with the goal instantiates \<open>?P\<close> to the  | 
|
| 69597 | 772  | 
function \<^term>\<open>\<lambda>n. \<Sum>{0..n::nat} = n*(n+1) div 2\<close>.  Now the proposition to
 | 
| 69505 | 773  | 
be proved in the base case can be written as \<open>?P 0\<close>, the induction  | 
774  | 
hypothesis as \<open>?P n\<close>, and the conclusion of the induction step as  | 
|
775  | 
\<open>?P(Suc n)\<close>.  | 
|
| 47269 | 776  | 
|
777  | 
Induction also provides the \isacom{case} idiom that abbreviates
 | 
|
778  | 
the \isacom{fix}-\isacom{assume} step. The above proof becomes
 | 
|
| 67406 | 779  | 
\<close>  | 
| 47269 | 780  | 
(*<*)lemma "\<Sum>{0..n::nat} = n*(n+1) div 2"(*>*)
 | 
781  | 
proof (induction n)  | 
|
782  | 
case 0  | 
|
783  | 
show ?case by simp  | 
|
784  | 
next  | 
|
785  | 
case (Suc n)  | 
|
786  | 
thus ?case by simp  | 
|
787  | 
qed  | 
|
788  | 
||
| 67406 | 789  | 
text\<open>  | 
| 69505 | 790  | 
The unknown \<open>?case\<close>\index{case?@\<open>?case\<close>|(} is set in each case to the required
 | 
791  | 
claim, i.e., \<open>?P 0\<close> and \mbox{\<open>?P(Suc n)\<close>} in the above proof,
 | 
|
792  | 
without requiring the user to define a \<open>?P\<close>. The general  | 
|
| 69597 | 793  | 
pattern for induction over \<^typ>\<open>nat\<close> is shown on the left-hand side:  | 
| 67406 | 794  | 
\<close>text_raw\<open>  | 
| 47269 | 795  | 
\begin{tabular}{@ {}ll@ {}}
 | 
796  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
797  | 
\isa{%
 | 
|
| 67406 | 798  | 
\<close>  | 
| 47269 | 799  | 
(*<*)lemma "P(n::nat)" proof -(*>*)  | 
800  | 
show "P(n)"  | 
|
801  | 
proof (induction n)  | 
|
802  | 
case 0  | 
|
| 67406 | 803  | 
  text_raw\<open>\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}\<close>
 | 
804  | 
show ?case (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 805  | 
next  | 
806  | 
case (Suc n)  | 
|
| 67406 | 807  | 
  text_raw\<open>\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}\<close>
 | 
808  | 
show ?case (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 809  | 
qed(*<*)qed(*>*)  | 
810  | 
||
| 67406 | 811  | 
text_raw \<open>}  | 
| 47269 | 812  | 
\end{minipage}
 | 
813  | 
&  | 
|
814  | 
\begin{minipage}[t]{.4\textwidth}
 | 
|
815  | 
~\\  | 
|
816  | 
~\\  | 
|
| 69505 | 817  | 
\isacom{let} \<open>?case = "P(0)"\<close>\\
 | 
| 47269 | 818  | 
~\\  | 
819  | 
~\\  | 
|
820  | 
~\\[1ex]  | 
|
| 69505 | 821  | 
\isacom{fix} \<open>n\<close> \isacom{assume} \<open>Suc: "P(n)"\<close>\\
 | 
822  | 
\isacom{let} \<open>?case = "P(Suc n)"\<close>\\
 | 
|
| 47269 | 823  | 
\end{minipage}
 | 
824  | 
\end{tabular}
 | 
|
825  | 
\medskip  | 
|
| 67406 | 826  | 
\<close>  | 
827  | 
text\<open>  | 
|
| 47269 | 828  | 
On the right side you can see what the \isacom{case} command
 | 
829  | 
on the left stands for.  | 
|
830  | 
||
831  | 
In case the goal is an implication, induction does one more thing: the  | 
|
832  | 
proposition to be proved in each case is not the whole implication but only  | 
|
833  | 
its conclusion; the premises of the implication are immediately made  | 
|
834  | 
assumptions of that case. That is, if in the above proof we replace  | 
|
| 69505 | 835  | 
\isacom{show}~\<open>"P(n)"\<close> by
 | 
836  | 
\mbox{\isacom{show}~\<open>"A(n) \<Longrightarrow> P(n)"\<close>}
 | 
|
837  | 
then \isacom{case}~\<open>0\<close> stands for
 | 
|
| 47269 | 838  | 
\begin{quote}
 | 
| 69505 | 839  | 
\isacom{assume} \ \<open>0: "A(0)"\<close>\\
 | 
840  | 
\isacom{let} \<open>?case = "P(0)"\<close>
 | 
|
| 47269 | 841  | 
\end{quote}
 | 
| 69505 | 842  | 
and \isacom{case}~\<open>(Suc n)\<close> stands for
 | 
| 47269 | 843  | 
\begin{quote}
 | 
| 69505 | 844  | 
\isacom{fix} \<open>n\<close>\\
 | 
845  | 
\isacom{assume} \<open>Suc:\<close>
 | 
|
846  | 
  \begin{tabular}[t]{l}\<open>"A(n) \<Longrightarrow> P(n)"\<close>\\\<open>"A(Suc n)"\<close>\end{tabular}\\
 | 
|
847  | 
\isacom{let} \<open>?case = "P(Suc n)"\<close>
 | 
|
| 47269 | 848  | 
\end{quote}
 | 
| 69505 | 849  | 
The list of assumptions \<open>Suc\<close> is actually subdivided  | 
850  | 
into \<open>Suc.IH\<close>, the induction hypotheses (here \<open>A(n) \<Longrightarrow> P(n)\<close>),  | 
|
851  | 
and \<open>Suc.prems\<close>, the premises of the goal being proved  | 
|
852  | 
(here \<open>A(Suc n)\<close>).  | 
|
| 47269 | 853  | 
|
854  | 
Induction works for any datatype.  | 
|
| 69505 | 855  | 
Proving a goal \<open>\<lbrakk> A\<^sub>1(x); \<dots>; A\<^sub>k(x) \<rbrakk> \<Longrightarrow> P(x)\<close>  | 
856  | 
by induction on \<open>x\<close> generates a proof obligation for each constructor  | 
|
857  | 
\<open>C\<close> of the datatype. The command \isacom{case}~\<open>(C x\<^sub>1 \<dots> x\<^sub>n)\<close>
 | 
|
| 47269 | 858  | 
performs the following steps:  | 
859  | 
\begin{enumerate}
 | 
|
| 69505 | 860  | 
\item \isacom{fix} \<open>x\<^sub>1 \<dots> x\<^sub>n\<close>
 | 
861  | 
\item \isacom{assume} the induction hypotheses (calling them \<open>C.IH\<close>\index{IH@\<open>.IH\<close>})
 | 
|
862  | 
 and the premises \mbox{\<open>A\<^sub>i(C x\<^sub>1 \<dots> x\<^sub>n)\<close>} (calling them \<open>C.prems\<close>\index{prems@\<open>.prems\<close>})
 | 
|
863  | 
and calling the whole list \<open>C\<close>  | 
|
864  | 
\item \isacom{let} \<open>?case = "P(C x\<^sub>1 \<dots> x\<^sub>n)"\<close>
 | 
|
| 47269 | 865  | 
\end{enumerate}
 | 
| 55361 | 866  | 
\index{structural induction|)}
 | 
| 47269 | 867  | 
|
| 65437 | 868  | 
|
869  | 
\ifsem\else  | 
|
870  | 
\subsection{Computation Induction}
 | 
|
871  | 
\index{rule induction}
 | 
|
872  | 
||
873  | 
In \autoref{sec:recursive-funs} we introduced computation induction and
 | 
|
874  | 
its realization in Isabelle: the definition  | 
|
875  | 
of a recursive function \<open>f\<close> via \isacom{fun} proves the corresponding computation
 | 
|
876  | 
induction rule called \<open>f.induct\<close>. Induction with this rule looks like in  | 
|
877  | 
\autoref{sec:recursive-funs}, but now with \isacom{proof} instead of \isacom{apply}:
 | 
|
878  | 
\begin{quote}
 | 
|
879  | 
\isacom{proof} (\<open>induction x\<^sub>1 \<dots> x\<^sub>k rule: f.induct\<close>)
 | 
|
880  | 
\end{quote}
 | 
|
881  | 
Just as for structural induction, this creates several cases, one for each  | 
|
882  | 
defining equation for \<open>f\<close>. By default (if the equations have not been named  | 
|
883  | 
by the user), the cases are numbered. That is, they are started by  | 
|
884  | 
\begin{quote}
 | 
|
885  | 
\isacom{case} (\<open>i x y ...\<close>)
 | 
|
886  | 
\end{quote}
 | 
|
887  | 
where \<open>i = 1,...,n\<close>, \<open>n\<close> is the number of equations defining \<open>f\<close>,  | 
|
888  | 
and \<open>x y ...\<close> are the variables in equation \<open>i\<close>. Note the following:  | 
|
889  | 
\begin{itemize}
 | 
|
890  | 
\item  | 
|
891  | 
Although \<open>i\<close> is an Isar name, \<open>i.IH\<close> (or similar) is not. You need  | 
|
892  | 
double quotes: "\<open>i.IH\<close>". When indexing the name, write "\<open>i.IH\<close>"(1),  | 
|
893  | 
not "\<open>i.IH\<close>(1)".  | 
|
894  | 
\item  | 
|
895  | 
If defining equations for \<open>f\<close> overlap, \isacom{fun} instantiates them to make
 | 
|
896  | 
them nonoverlapping. This means that one user-provided equation may lead to  | 
|
897  | 
several equations and thus to several cases in the induction rule.  | 
|
898  | 
These have names of the form "\<open>i_j\<close>", where \<open>i\<close> is the number of the original  | 
|
899  | 
equation and the system-generated \<open>j\<close> indicates the subcase.  | 
|
900  | 
\end{itemize}
 | 
|
901  | 
In Isabelle/jEdit, the \<open>induction\<close> proof method displays a proof skeleton  | 
|
902  | 
with all \isacom{case}s. This is particularly useful for computation induction
 | 
|
903  | 
and the following rule induction.  | 
|
904  | 
\fi  | 
|
905  | 
||
906  | 
||
| 52361 | 907  | 
\subsection{Rule Induction}
 | 
| 55361 | 908  | 
\index{rule induction|(}
 | 
| 47269 | 909  | 
|
910  | 
Recall the inductive and recursive definitions of even numbers in  | 
|
911  | 
\autoref{sec:inductive-defs}:
 | 
|
| 67406 | 912  | 
\<close>  | 
| 47269 | 913  | 
|
914  | 
inductive ev :: "nat \<Rightarrow> bool" where  | 
|
915  | 
ev0: "ev 0" |  | 
|
916  | 
evSS: "ev n \<Longrightarrow> ev(Suc(Suc n))"  | 
|
917  | 
||
| 61429 | 918  | 
fun evn :: "nat \<Rightarrow> bool" where  | 
919  | 
"evn 0 = True" |  | 
|
920  | 
"evn (Suc 0) = False" |  | 
|
921  | 
"evn (Suc(Suc n)) = evn n"  | 
|
| 47269 | 922  | 
|
| 69597 | 923  | 
text\<open>We recast the proof of \<^prop>\<open>ev n \<Longrightarrow> evn n\<close> in Isar. The  | 
| 47269 | 924  | 
left column shows the actual proof text, the right column shows  | 
| 67406 | 925  | 
the implicit effect of the two \isacom{case} commands:\<close>text_raw\<open>
 | 
| 47269 | 926  | 
\begin{tabular}{@ {}l@ {\qquad}l@ {}}
 | 
927  | 
\begin{minipage}[t]{.5\textwidth}
 | 
|
928  | 
\isa{%
 | 
|
| 67406 | 929  | 
\<close>  | 
| 47269 | 930  | 
|
| 61429 | 931  | 
lemma "ev n \<Longrightarrow> evn n"  | 
| 47269 | 932  | 
proof(induction rule: ev.induct)  | 
933  | 
case ev0  | 
|
934  | 
show ?case by simp  | 
|
935  | 
next  | 
|
936  | 
case evSS  | 
|
937  | 
||
938  | 
||
939  | 
||
940  | 
thus ?case by simp  | 
|
941  | 
qed  | 
|
942  | 
||
| 67406 | 943  | 
text_raw \<open>}  | 
| 47269 | 944  | 
\end{minipage}
 | 
945  | 
&  | 
|
946  | 
\begin{minipage}[t]{.5\textwidth}
 | 
|
947  | 
~\\  | 
|
948  | 
~\\  | 
|
| 69505 | 949  | 
\isacom{let} \<open>?case = "evn 0"\<close>\\
 | 
| 47269 | 950  | 
~\\  | 
951  | 
~\\  | 
|
| 69505 | 952  | 
\isacom{fix} \<open>n\<close>\\
 | 
953  | 
\isacom{assume} \<open>evSS:\<close>
 | 
|
954  | 
  \begin{tabular}[t]{l} \<open>"ev n"\<close>\\\<open>"evn n"\<close>\end{tabular}\\
 | 
|
955  | 
\isacom{let} \<open>?case = "evn(Suc(Suc n))"\<close>\\
 | 
|
| 47269 | 956  | 
\end{minipage}
 | 
957  | 
\end{tabular}
 | 
|
958  | 
\medskip  | 
|
| 67406 | 959  | 
\<close>  | 
960  | 
text\<open>  | 
|
| 47269 | 961  | 
The proof resembles structural induction, but the induction rule is given  | 
962  | 
explicitly and the names of the cases are the names of the rules in the  | 
|
963  | 
inductive definition.  | 
|
964  | 
Let us examine the two assumptions named @{thm[source]evSS}:
 | 
|
| 69597 | 965  | 
\<^prop>\<open>ev n\<close> is the premise of rule @{thm[source]evSS}, which we may assume
 | 
966  | 
because we are in the case where that rule was used; \<^prop>\<open>evn n\<close>  | 
|
| 47269 | 967  | 
is the induction hypothesis.  | 
968  | 
\begin{warn}
 | 
|
969  | 
Because each \isacom{case} command introduces a list of assumptions
 | 
|
970  | 
named like the case name, which is the name of a rule of the inductive  | 
|
971  | 
definition, those rules now need to be accessed with a qualified name, here  | 
|
| 58522 | 972  | 
@{thm[source] ev.ev0} and @{thm[source] ev.evSS}.
 | 
| 47269 | 973  | 
\end{warn}
 | 
974  | 
||
975  | 
In the case @{thm[source]evSS} of the proof above we have pretended that the
 | 
|
| 69505 | 976  | 
system fixes a variable \<open>n\<close>. But unless the user provides the name  | 
977  | 
\<open>n\<close>, the system will just invent its own name that cannot be referred  | 
|
| 47269 | 978  | 
to. In the above proof, we do not need to refer to it, hence we do not give  | 
979  | 
it a specific name. In case one needs to refer to it one writes  | 
|
980  | 
\begin{quote}
 | 
|
| 69505 | 981  | 
\isacom{case} \<open>(evSS m)\<close>
 | 
| 47269 | 982  | 
\end{quote}
 | 
| 69505 | 983  | 
like \isacom{case}~\<open>(Suc n)\<close> in earlier structural inductions.
 | 
984  | 
The name \<open>m\<close> is an arbitrary choice. As a result,  | 
|
| 47269 | 985  | 
case @{thm[source] evSS} is derived from a renamed version of
 | 
| 69505 | 986  | 
rule @{thm[source] evSS}: \<open>ev m \<Longrightarrow> ev(Suc(Suc m))\<close>.
 | 
987  | 
Here is an example with a (contrived) intermediate step that refers to \<open>m\<close>:  | 
|
| 67406 | 988  | 
\<close>  | 
| 47269 | 989  | 
|
| 61429 | 990  | 
lemma "ev n \<Longrightarrow> evn n"  | 
| 47269 | 991  | 
proof(induction rule: ev.induct)  | 
992  | 
case ev0 show ?case by simp  | 
|
993  | 
next  | 
|
994  | 
case (evSS m)  | 
|
| 61429 | 995  | 
have "evn(Suc(Suc m)) = evn m" by simp  | 
| 67406 | 996  | 
thus ?case using \<open>evn m\<close> by blast  | 
| 47269 | 997  | 
qed  | 
998  | 
||
| 67406 | 999  | 
text\<open>  | 
| 47269 | 1000  | 
\indent  | 
| 69505 | 1001  | 
In general, let \<open>I\<close> be a (for simplicity unary) inductively defined  | 
1002  | 
predicate and let the rules in the definition of \<open>I\<close>  | 
|
1003  | 
be called \<open>rule\<^sub>1\<close>, \dots, \<open>rule\<^sub>n\<close>. A proof by rule  | 
|
1004  | 
induction follows this pattern:\index{inductionrule@\<open>induction ... rule:\<close>}
 | 
|
| 67406 | 1005  | 
\<close>  | 
| 47269 | 1006  | 
|
1007  | 
(*<*)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52718 
diff
changeset
 | 
1008  | 
inductive I where rule\<^sub>1: "I()" | rule\<^sub>2: "I()" | rule\<^sub>n: "I()"  | 
| 47269 | 1009  | 
lemma "I x \<Longrightarrow> P x" proof-(*>*)  | 
1010  | 
show "I x \<Longrightarrow> P x"  | 
|
1011  | 
proof(induction rule: I.induct)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52718 
diff
changeset
 | 
1012  | 
case rule\<^sub>1  | 
| 67406 | 1013  | 
  text_raw\<open>\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\<close>
 | 
1014  | 
show ?case (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 1015  | 
next  | 
| 67406 | 1016  | 
  text_raw\<open>\\[-.4ex]$\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\<close>
 | 
| 47269 | 1017  | 
(*<*)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52718 
diff
changeset
 | 
1018  | 
case rule\<^sub>2  | 
| 47269 | 1019  | 
show ?case sorry  | 
1020  | 
(*>*)  | 
|
1021  | 
next  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52718 
diff
changeset
 | 
1022  | 
case rule\<^sub>n  | 
| 67406 | 1023  | 
  text_raw\<open>\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\<close>
 | 
1024  | 
show ?case (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>  | 
|
| 47269 | 1025  | 
qed(*<*)qed(*>*)  | 
1026  | 
||
| 67406 | 1027  | 
text\<open>  | 
| 47269 | 1028  | 
One can provide explicit variable names by writing  | 
| 69505 | 1029  | 
\isacom{case}~\<open>(rule\<^sub>i x\<^sub>1 \<dots> x\<^sub>k)\<close>, thus renaming the first \<open>k\<close>
 | 
1030  | 
free variables in rule \<open>i\<close> to \<open>x\<^sub>1 \<dots> x\<^sub>k\<close>,  | 
|
1031  | 
going through rule \<open>i\<close> from left to right.  | 
|
| 47269 | 1032  | 
|
| 52361 | 1033  | 
\subsection{Assumption Naming}
 | 
| 51443 | 1034  | 
\label{sec:assm-naming}
 | 
| 47269 | 1035  | 
|
| 69505 | 1036  | 
In any induction, \isacom{case}~\<open>name\<close> sets up a list of assumptions
 | 
1037  | 
also called \<open>name\<close>, which is subdivided into three parts:  | 
|
| 47269 | 1038  | 
\begin{description}
 | 
| 69505 | 1039  | 
\item[\<open>name.IH\<close>]\index{IH@\<open>.IH\<close>} contains the induction hypotheses.
 | 
1040  | 
\item[\<open>name.hyps\<close>]\index{hyps@\<open>.hyps\<close>} contains all the other hypotheses of this case in the
 | 
|
| 47269 | 1041  | 
induction rule. For rule inductions these are the hypotheses of rule  | 
| 69505 | 1042  | 
\<open>name\<close>, for structural inductions these are empty.  | 
1043  | 
\item[\<open>name.prems\<close>]\index{prems@\<open>.prems\<close>} contains the (suitably instantiated) premises
 | 
|
1044  | 
of the statement being proved, i.e., the \<open>A\<^sub>i\<close> when  | 
|
1045  | 
proving \<open>\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A\<close>.  | 
|
| 47269 | 1046  | 
\end{description}
 | 
1047  | 
\begin{warn}
 | 
|
| 69505 | 1048  | 
Proof method \<open>induct\<close> differs from \<open>induction\<close>  | 
1049  | 
only in this naming policy: \<open>induct\<close> does not distinguish  | 
|
1050  | 
\<open>IH\<close> from \<open>hyps\<close> but subsumes \<open>IH\<close> under \<open>hyps\<close>.  | 
|
| 47269 | 1051  | 
\end{warn}
 | 
1052  | 
||
1053  | 
More complicated inductive proofs than the ones we have seen so far  | 
|
| 69505 | 1054  | 
often need to refer to specific assumptions --- just \<open>name\<close> or even  | 
1055  | 
\<open>name.prems\<close> and \<open>name.IH\<close> can be too unspecific.  | 
|
| 58504 | 1056  | 
This is where the indexing of fact lists comes in handy, e.g.,  | 
| 69505 | 1057  | 
\<open>name.IH(2)\<close> or \<open>name.prems(1-2)\<close>.  | 
| 47269 | 1058  | 
|
| 52361 | 1059  | 
\subsection{Rule Inversion}
 | 
1060  | 
\label{sec:rule-inversion}
 | 
|
| 55361 | 1061  | 
\index{rule inversion|(}
 | 
| 47269 | 1062  | 
|
| 47711 | 1063  | 
Rule inversion is case analysis of which rule could have been used to  | 
| 55361 | 1064  | 
derive some fact. The name \conceptnoidx{rule inversion} emphasizes that we are
 | 
| 47269 | 1065  | 
reasoning backwards: by which rules could some given fact have been proved?  | 
| 69597 | 1066  | 
For the inductive definition of \<^const>\<open>ev\<close>, rule inversion can be summarized  | 
| 47269 | 1067  | 
like this:  | 
| 67613 | 1068  | 
@{prop[display]"ev n \<Longrightarrow> n = 0 \<or> (\<exists>k. n = Suc(Suc k) \<and> ev k)"}
 | 
| 47711 | 1069  | 
The realisation in Isabelle is a case analysis.  | 
| 69597 | 1070  | 
A simple example is the proof that \<^prop>\<open>ev n \<Longrightarrow> ev (n - 2)\<close>. We  | 
| 47269 | 1071  | 
already went through the details informally in \autoref{sec:Logic:even}. This
 | 
1072  | 
is the Isar proof:  | 
|
| 67406 | 1073  | 
\<close>  | 
| 47269 | 1074  | 
(*<*)  | 
1075  | 
notepad  | 
|
1076  | 
begin fix n  | 
|
1077  | 
(*>*)  | 
|
1078  | 
assume "ev n"  | 
|
1079  | 
from this have "ev(n - 2)"  | 
|
1080  | 
proof cases  | 
|
1081  | 
case ev0 thus "ev(n - 2)" by (simp add: ev.ev0)  | 
|
1082  | 
next  | 
|
1083  | 
case (evSS k) thus "ev(n - 2)" by (simp add: ev.evSS)  | 
|
1084  | 
qed  | 
|
1085  | 
(*<*)  | 
|
1086  | 
end  | 
|
1087  | 
(*>*)  | 
|
1088  | 
||
| 67406 | 1089  | 
text\<open>The key point here is that a case analysis over some inductively  | 
| 47269 | 1090  | 
defined predicate is triggered by piping the given fact  | 
| 69505 | 1091  | 
(here: \isacom{from}~\<open>this\<close>) into a proof by \<open>cases\<close>.
 | 
1092  | 
Let us examine the assumptions available in each case. In case \<open>ev0\<close>  | 
|
| 69597 | 1093  | 
we have \<open>n = 0\<close> and in case \<open>evSS\<close> we have \<^prop>\<open>n = Suc(Suc k)\<close>  | 
1094  | 
and \<^prop>\<open>ev k\<close>. In each case the assumptions are available under the name  | 
|
| 56989 | 1095  | 
of the case; there is no fine-grained naming schema like there is for induction.  | 
| 47269 | 1096  | 
|
| 47704 | 1097  | 
Sometimes some rules could not have been used to derive the given fact  | 
| 47269 | 1098  | 
because constructors clash. As an extreme example consider  | 
| 69597 | 1099  | 
rule inversion applied to \<^prop>\<open>ev(Suc 0)\<close>: neither rule \<open>ev0\<close> nor  | 
1100  | 
rule \<open>evSS\<close> can yield \<^prop>\<open>ev(Suc 0)\<close> because \<open>Suc 0\<close> unifies  | 
|
1101  | 
neither with \<open>0\<close> nor with \<^term>\<open>Suc(Suc n)\<close>. Impossible cases do not  | 
|
1102  | 
have to be proved. Hence we can prove anything from \<^prop>\<open>ev(Suc 0)\<close>:  | 
|
| 67406 | 1103  | 
\<close>  | 
| 47269 | 1104  | 
(*<*)  | 
1105  | 
notepad begin fix P  | 
|
1106  | 
(*>*)  | 
|
1107  | 
assume "ev(Suc 0)" then have P by cases  | 
|
1108  | 
(*<*)  | 
|
1109  | 
end  | 
|
1110  | 
(*>*)  | 
|
1111  | 
||
| 69597 | 1112  | 
text\<open>That is, \<^prop>\<open>ev(Suc 0)\<close> is simply not provable:\<close>  | 
| 47269 | 1113  | 
|
1114  | 
lemma "\<not> ev(Suc 0)"  | 
|
1115  | 
proof  | 
|
1116  | 
assume "ev(Suc 0)" then show False by cases  | 
|
1117  | 
qed  | 
|
1118  | 
||
| 67406 | 1119  | 
text\<open>Normally not all cases will be impossible. As a simple exercise,  | 
| 69597 | 1120  | 
prove that \mbox{\<^prop>\<open>\<not> ev(Suc(Suc(Suc 0)))\<close>.}
 | 
| 51443 | 1121  | 
|
| 52361 | 1122  | 
\subsection{Advanced Rule Induction}
 | 
| 51445 | 1123  | 
\label{sec:advanced-rule-induction}
 | 
| 51443 | 1124  | 
|
| 69505 | 1125  | 
So far, rule induction was always applied to goals of the form \<open>I x y z \<Longrightarrow> \<dots>\<close>  | 
1126  | 
where \<open>I\<close> is some inductively defined predicate and \<open>x\<close>, \<open>y\<close>, \<open>z\<close>  | 
|
| 51443 | 1127  | 
are variables. In some rare situations one needs to deal with an assumption where  | 
| 69505 | 1128  | 
not all arguments \<open>r\<close>, \<open>s\<close>, \<open>t\<close> are variables:  | 
| 51443 | 1129  | 
\begin{isabelle}
 | 
| 69505 | 1130  | 
\isacom{lemma} \<open>"I r s t \<Longrightarrow> \<dots>"\<close>
 | 
| 51443 | 1131  | 
\end{isabelle}
 | 
1132  | 
Applying the standard form of  | 
|
| 54577 | 1133  | 
rule induction in such a situation will lead to strange and typically unprovable goals.  | 
| 51443 | 1134  | 
We can easily reduce this situation to the standard one by introducing  | 
| 69505 | 1135  | 
new variables \<open>x\<close>, \<open>y\<close>, \<open>z\<close> and reformulating the goal like this:  | 
| 51443 | 1136  | 
\begin{isabelle}
 | 
| 69505 | 1137  | 
\isacom{lemma} \<open>"I x y z \<Longrightarrow> x = r \<Longrightarrow> y = s \<Longrightarrow> z = t \<Longrightarrow> \<dots>"\<close>
 | 
| 51443 | 1138  | 
\end{isabelle}
 | 
| 56989 | 1139  | 
Standard rule induction will work fine now, provided the free variables in  | 
| 69505 | 1140  | 
\<open>r\<close>, \<open>s\<close>, \<open>t\<close> are generalized via \<open>arbitrary\<close>.  | 
| 51443 | 1141  | 
|
1142  | 
However, induction can do the above transformation for us, behind the curtains, so we never  | 
|
1143  | 
need to see the expanded version of the lemma. This is what we need to write:  | 
|
1144  | 
\begin{isabelle}
 | 
|
| 69505 | 1145  | 
\isacom{lemma} \<open>"I r s t \<Longrightarrow> \<dots>"\<close>\isanewline
 | 
1146  | 
\isacom{proof}\<open>(induction "r" "s" "t" arbitrary: \<dots> rule: I.induct)\<close>\index{inductionrule@\<open>induction ... rule:\<close>}\index{arbitrary@\<open>arbitrary:\<close>}
 | 
|
| 51443 | 1147  | 
\end{isabelle}
 | 
| 58605 | 1148  | 
Like for rule inversion, cases that are impossible because of constructor clashes  | 
| 67406 | 1149  | 
will not show up at all. Here is a concrete example:\<close>  | 
| 51443 | 1150  | 
|
1151  | 
lemma "ev (Suc m) \<Longrightarrow> \<not> ev m"  | 
|
1152  | 
proof(induction "Suc m" arbitrary: m rule: ev.induct)  | 
|
1153  | 
fix n assume IH: "\<And>m. n = Suc m \<Longrightarrow> \<not> ev m"  | 
|
1154  | 
show "\<not> ev (Suc n)"  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
1155  | 
proof \<comment> \<open>contradiction\<close>  | 
| 51443 | 1156  | 
assume "ev(Suc n)"  | 
1157  | 
thus False  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
1158  | 
proof cases \<comment> \<open>rule inversion\<close>  | 
| 51443 | 1159  | 
fix k assume "n = Suc k" "ev k"  | 
1160  | 
thus False using IH by auto  | 
|
1161  | 
qed  | 
|
1162  | 
qed  | 
|
1163  | 
qed  | 
|
1164  | 
||
| 67406 | 1165  | 
text\<open>  | 
| 51443 | 1166  | 
Remarks:  | 
1167  | 
\begin{itemize}
 | 
|
1168  | 
\item  | 
|
| 69505 | 1169  | 
Instead of the \isacom{case} and \<open>?case\<close> magic we have spelled all formulas out.
 | 
| 51443 | 1170  | 
This is merely for greater clarity.  | 
1171  | 
\item  | 
|
1172  | 
We only need to deal with one case because the @{thm[source] ev0} case is impossible.
 | 
|
1173  | 
\item  | 
|
| 69505 | 1174  | 
The form of the \<open>IH\<close> shows us that internally the lemma was expanded as explained  | 
| 51443 | 1175  | 
above: \noquotes{@{prop[source]"ev x \<Longrightarrow> x = Suc m \<Longrightarrow> \<not> ev m"}}.
 | 
1176  | 
\item  | 
|
| 69597 | 1177  | 
The goal \<^prop>\<open>\<not> ev (Suc n)\<close> may surprise. The expanded version of the lemma  | 
1178  | 
would suggest that we have a \isacom{fix} \<open>m\<close> \isacom{assume} \<^prop>\<open>Suc(Suc n) = Suc m\<close>
 | 
|
1179  | 
and need to show \<^prop>\<open>\<not> ev m\<close>. What happened is that Isabelle immediately  | 
|
1180  | 
simplified \<^prop>\<open>Suc(Suc n) = Suc m\<close> to \<^prop>\<open>Suc n = m\<close> and could then eliminate  | 
|
| 69505 | 1181  | 
\<open>m\<close>. Beware of such nice surprises with this advanced form of induction.  | 
| 51443 | 1182  | 
\end{itemize}
 | 
1183  | 
\begin{warn}
 | 
|
| 69505 | 1184  | 
This advanced form of induction does not support the \<open>IH\<close>  | 
| 51443 | 1185  | 
naming schema explained in \autoref{sec:assm-naming}:
 | 
| 69505 | 1186  | 
the induction hypotheses are instead found under the name \<open>hyps\<close>,  | 
| 56989 | 1187  | 
as they are for the simpler  | 
| 69505 | 1188  | 
\<open>induct\<close> method.  | 
| 51443 | 1189  | 
\end{warn}
 | 
| 55361 | 1190  | 
\index{induction|)}
 | 
| 69505 | 1191  | 
\index{cases@\<open>cases\<close>|)}
 | 
| 55361 | 1192  | 
\index{case@\isacom{case}|)}
 | 
| 69505 | 1193  | 
\index{case?@\<open>?case\<close>|)}
 | 
| 55361 | 1194  | 
\index{rule induction|)}
 | 
1195  | 
\index{rule inversion|)}
 | 
|
| 54218 | 1196  | 
|
| 54436 | 1197  | 
\subsection*{Exercises}
 | 
| 52593 | 1198  | 
|
| 54232 | 1199  | 
|
1200  | 
\exercise  | 
|
| 54292 | 1201  | 
Give a structured proof by rule inversion:  | 
| 67406 | 1202  | 
\<close>  | 
| 54232 | 1203  | 
|
1204  | 
lemma assumes a: "ev(Suc(Suc n))" shows "ev n"  | 
|
1205  | 
(*<*)oops(*>*)  | 
|
1206  | 
||
| 67406 | 1207  | 
text\<open>  | 
| 54232 | 1208  | 
\endexercise  | 
1209  | 
||
| 52593 | 1210  | 
\begin{exercise}
 | 
| 69597 | 1211  | 
Give a structured proof of \<^prop>\<open>\<not> ev(Suc(Suc(Suc 0)))\<close>  | 
| 54232 | 1212  | 
by rule inversions. If there are no cases to be proved you can close  | 
| 61013 | 1213  | 
a proof immediately with \isacom{qed}.
 | 
| 54218 | 1214  | 
\end{exercise}
 | 
1215  | 
||
1216  | 
\begin{exercise}
 | 
|
| 69505 | 1217  | 
Recall predicate \<open>star\<close> from \autoref{sec:star} and \<open>iter\<close>
 | 
| 69597 | 1218  | 
from Exercise~\ref{exe:iter}. Prove \<^prop>\<open>iter r n x y \<Longrightarrow> star r x y\<close>
 | 
| 56989 | 1219  | 
in a structured style; do not just sledgehammer each case of the  | 
| 54292 | 1220  | 
required induction.  | 
1221  | 
\end{exercise}
 | 
|
1222  | 
||
1223  | 
\begin{exercise}
 | 
|
| 69597 | 1224  | 
Define a recursive function \<open>elems ::\<close> \<^typ>\<open>'a list \<Rightarrow> 'a set\<close>  | 
1225  | 
and prove \<^prop>\<open>x \<in> elems xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> elems ys\<close>.  | 
|
| 52593 | 1226  | 
\end{exercise}
 | 
| 61012 | 1227  | 
|
1228  | 
\begin{exercise}
 | 
|
1229  | 
Extend Exercise~\ref{exe:cfg} with a function that checks if some
 | 
|
| 69505 | 1230  | 
\mbox{\<open>alpha list\<close>} is a balanced
 | 
| 61022 | 1231  | 
string of parentheses. More precisely, define a \mbox{recursive} function
 | 
| 69597 | 1232  | 
\<open>balanced :: nat \<Rightarrow> alpha list \<Rightarrow> bool\<close> such that \<^term>\<open>balanced n w\<close>  | 
| 69505 | 1233  | 
is true iff (informally) \<open>S (a\<^sup>n @ w)\<close>. Formally, prove that  | 
| 69597 | 1234  | 
\<^prop>\<open>balanced n w \<longleftrightarrow> S (replicate n a @ w)\<close> where  | 
1235  | 
\<^const>\<open>replicate\<close> \<open>::\<close> \<^typ>\<open>nat \<Rightarrow> 'a \<Rightarrow> 'a list\<close> is predefined  | 
|
1236  | 
and \<^term>\<open>replicate n x\<close> yields the list \<open>[x, \<dots>, x]\<close> of length \<open>n\<close>.  | 
|
| 61012 | 1237  | 
\end{exercise}
 | 
| 67406 | 1238  | 
\<close>  | 
| 47269 | 1239  | 
|
1240  | 
(*<*)  | 
|
1241  | 
end  | 
|
1242  | 
(*>*)  |