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