| 5375 |      1 | \chapter{Basic Concepts}
 | 
|  |      2 | 
 | 
|  |      3 | \section{Introduction}
 | 
|  |      4 | 
 | 
|  |      5 | This is a tutorial on how to use Isabelle/HOL as a specification and
 | 
|  |      6 | verification system. Isabelle is a generic system for implementing logical
 | 
|  |      7 | formalisms, and Isabelle/HOL is the specialization of Isabelle for
 | 
|  |      8 | HOL, which abbreviates Higher-Order Logic. We introduce HOL step by step
 | 
|  |      9 | following the equation
 | 
|  |     10 | \[ \mbox{HOL} = \mbox{Functional Programming} + \mbox{Logic}. \]
 | 
|  |     11 | We assume that the reader is familiar with the basic concepts of both fields.
 | 
|  |     12 | For excellent introductions to functional programming consult the textbooks
 | 
| 6606 |     13 | by Bird and Wadler~\cite{Bird-Wadler} or Paulson~\cite{paulson-ml2}.  Although
 | 
| 5375 |     14 | this tutorial initially concentrates on functional programming, do not be
 | 
|  |     15 | misled: HOL can express most mathematical concepts, and functional
 | 
|  |     16 | programming is just one particularly simple and ubiquitous instance.
 | 
|  |     17 | 
 | 
|  |     18 | A tutorial is by definition incomplete. To fully exploit the power of the
 | 
| 6606 |     19 | system you need to consult the Isabelle Reference Manual~\cite{isabelle-ref}
 | 
|  |     20 | for details about Isabelle and the Isabelle/HOL manual~\cite{isabelle-HOL}
 | 
|  |     21 | for details relating to HOL. Both manuals have a comprehensive index.
 | 
| 5375 |     22 | 
 | 
|  |     23 | \section{Theories, proofs and interaction}
 | 
|  |     24 | \label{sec:Basic:Theories}
 | 
|  |     25 | 
 | 
|  |     26 | Working with Isabelle means creating two different kinds of documents:
 | 
|  |     27 | theories and proof scripts. Roughly speaking, a \bfindex{theory} is a named
 | 
|  |     28 | collection of types and functions, much like a module in a programming
 | 
|  |     29 | language or a specification in a specification language. In fact, theories in
 | 
|  |     30 | HOL can be either. Theories must reside in files with the suffix
 | 
|  |     31 | \texttt{.thy}. The general format of a theory file \texttt{T.thy} is
 | 
|  |     32 | \begin{ttbox}
 | 
|  |     33 | T = B\(@1\) + \(\cdots\) + B\(@n\) +
 | 
| 6584 |     34 | \({\langle}declarations{\rangle}\)
 | 
| 5375 |     35 | end
 | 
|  |     36 | \end{ttbox}
 | 
|  |     37 | where \texttt{B}$@1$, \dots, \texttt{B}$@n$ are the names of existing
 | 
| 6584 |     38 | theories that \texttt{T} is based on and ${\langle}declarations{\rangle}$ stands for the
 | 
| 5375 |     39 | newly introduced concepts (types, functions etc). The \texttt{B}$@i$ are the
 | 
|  |     40 | direct \textbf{parent theories}\indexbold{parent theory} of \texttt{T}.
 | 
|  |     41 | Everything defined in the parent theories (and their parents \dots) is
 | 
|  |     42 | automatically visible. To avoid name clashes, identifiers can be qualified by
 | 
|  |     43 | theory names as in \texttt{T.f} and \texttt{B.f}. HOL's theory library is
 | 
| 6148 |     44 | available online at
 | 
| 6628 |     45 | 
 | 
|  |     46 | \begin{center}\small
 | 
| 9255 |     47 |     \url{http://isabelle.in.tum.de/library/}
 | 
| 6628 |     48 | \end{center}
 | 
|  |     49 | 
 | 
| 6148 |     50 | and is recommended browsing.
 | 
| 5375 |     51 | \begin{warn}
 | 
|  |     52 |   HOL contains a theory \ttindexbold{Main}, the union of all the basic
 | 
|  |     53 |   predefined theories like arithmetic, lists, sets, etc.\ (see the online
 | 
|  |     54 |   library).  Unless you know what you are doing, always include \texttt{Main}
 | 
|  |     55 |   as a direct or indirect parent theory of all your theories.
 | 
|  |     56 | \end{warn}
 | 
|  |     57 | 
 | 
|  |     58 | This tutorial is concerned with introducing you to the different linguistic
 | 
| 6584 |     59 | constructs that can fill ${\langle}declarations{\rangle}$ in the above theory template.
 | 
| 5375 |     60 | A complete grammar of the basic constructs is found in Appendix~A
 | 
| 6606 |     61 | of~\cite{isabelle-ref}, for reference in times of doubt.
 | 
| 5375 |     62 | 
 | 
|  |     63 | The tutorial is also concerned with showing you how to prove theorems about
 | 
|  |     64 | the concepts in a theory. This involves invoking predefined theorem proving
 | 
|  |     65 | commands. Because Isabelle is written in the programming language
 | 
|  |     66 | ML,\footnote{Many concepts in HOL and ML are similar. Make sure you do not
 | 
|  |     67 |   confuse the two levels.} interacting with Isabelle means calling ML
 | 
|  |     68 | functions. Hence \bfindex{proof scripts} are sequences of calls to ML
 | 
|  |     69 | functions that perform specific theorem proving tasks. Nevertheless,
 | 
|  |     70 | familiarity with ML is absolutely not required.  All proof scripts for theory
 | 
|  |     71 | \texttt{T} (defined in file \texttt{T.thy}) should be contained in file
 | 
|  |     72 | \texttt{T.ML}. Theory and proof scripts are loaded (and checked!) by calling
 | 
|  |     73 | the ML function \ttindexbold{use_thy}:
 | 
|  |     74 | \begin{ttbox}
 | 
|  |     75 | use_thy "T";
 | 
|  |     76 | \end{ttbox}
 | 
|  |     77 | 
 | 
|  |     78 | There are more advanced interfaces for Isabelle that hide the ML level from
 | 
|  |     79 | you and replace function calls by menu selection. There is even a special
 | 
|  |     80 | font with mathematical symbols. For details see the Isabelle home page.  This
 | 
|  |     81 | tutorial concentrates on the bare essentials and ignores such niceties.
 | 
|  |     82 | 
 | 
|  |     83 | \section{Types, terms and formulae}
 | 
|  |     84 | \label{sec:TypesTermsForms}
 | 
|  |     85 | 
 | 
|  |     86 | Embedded in the declarations of a theory are the types, terms and formulae of
 | 
|  |     87 | HOL. HOL is a typed logic whose type system resembles that of functional
 | 
|  |     88 | programming languages like ML or Haskell. Thus there are
 | 
|  |     89 | \begin{description}
 | 
|  |     90 | \item[base types,] in particular \ttindex{bool}, the type of truth values,
 | 
|  |     91 | and \ttindex{nat}, the type of natural numbers.
 | 
|  |     92 | \item[type constructors,] in particular \ttindex{list}, the type of
 | 
|  |     93 | lists, and \ttindex{set}, the type of sets. Type constructors are written
 | 
|  |     94 | postfix, e.g.\ \texttt{(nat)list} is the type of lists whose elements are
 | 
|  |     95 | natural numbers. Parentheses around single arguments can be dropped (as in
 | 
|  |     96 | \texttt{nat list}), multiple arguments are separated by commas (as in
 | 
|  |     97 | \texttt{(bool,nat)foo}).
 | 
|  |     98 | \item[function types,] denoted by \ttindexbold{=>}. In HOL
 | 
|  |     99 | \texttt{=>} represents {\em total} functions only. As is customary,
 | 
|  |    100 | \texttt{$\tau@1$ => $\tau@2$ => $\tau@3$} means
 | 
|  |    101 | \texttt{$\tau@1$ => ($\tau@2$ => $\tau@3$)}. Isabelle also supports the
 | 
|  |    102 | notation \texttt{[$\tau@1,\dots,\tau@n$] => $\tau$} which abbreviates
 | 
|  |    103 | \texttt{$\tau@1$ => $\cdots$ => $\tau@n$ => $\tau$}.
 | 
|  |    104 | \item[type variables,] denoted by \texttt{'a}, \texttt{'b} etc, just like in
 | 
|  |    105 | ML. They give rise to polymorphic types like \texttt{'a => 'a}, the type of the
 | 
|  |    106 | identity function.
 | 
|  |    107 | \end{description}
 | 
|  |    108 | \begin{warn}
 | 
|  |    109 |   Types are extremely important because they prevent us from writing
 | 
|  |    110 |   nonsense.  Isabelle insists that all terms and formulae must be well-typed
 | 
|  |    111 |   and will print an error message if a type mismatch is encountered. To
 | 
|  |    112 |   reduce the amount of explicit type information that needs to be provided by
 | 
|  |    113 |   the user, Isabelle infers the type of all variables automatically (this is
 | 
|  |    114 |   called \bfindex{type inference}) and keeps quiet about it. Occasionally
 | 
|  |    115 |   this may lead to misunderstandings between you and the system. If anything
 | 
|  |    116 |   strange happens, we recommend to set the flag \ttindexbold{show_types} that
 | 
|  |    117 |   tells Isabelle to display type information that is usually suppressed:
 | 
|  |    118 |   simply type
 | 
|  |    119 | \begin{ttbox}
 | 
|  |    120 | set show_types;
 | 
|  |    121 | \end{ttbox}
 | 
|  |    122 | 
 | 
|  |    123 | \noindent
 | 
|  |    124 | at the ML-level. This can be reversed by \texttt{reset show_types;}.
 | 
|  |    125 | \end{warn}
 | 
|  |    126 | 
 | 
|  |    127 | 
 | 
|  |    128 | \textbf{Terms}\indexbold{term}
 | 
|  |    129 | are formed as in functional programming by applying functions to
 | 
|  |    130 | arguments. If \texttt{f} is a function of type \texttt{$\tau@1$ => $\tau@2$}
 | 
|  |    131 | and \texttt{t} is a term of type $\tau@1$ then \texttt{f~t} is a term of type
 | 
|  |    132 | $\tau@2$. HOL also supports infix functions like \texttt{+} and some basic
 | 
|  |    133 | constructs from functional programming:
 | 
|  |    134 | \begin{description}
 | 
|  |    135 | \item[\texttt{if $b$ then $t@1$ else $t@2$}]\indexbold{*if}
 | 
|  |    136 | means what you think it means and requires that
 | 
|  |    137 | $b$ is of type \texttt{bool} and $t@1$ and $t@2$ are of the same type.
 | 
|  |    138 | \item[\texttt{let $x$ = $t$ in $u$}]\indexbold{*let}
 | 
|  |    139 | is equivalent to $u$ where all occurrences of $x$ have been replaced by
 | 
|  |    140 | $t$. For example,
 | 
|  |    141 | \texttt{let x = 0 in x+x} means \texttt{0+0}. Multiple bindings are separated
 | 
|  |    142 | by semicolons: \texttt{let $x@1$ = $t@1$; \dots; $x@n$ = $t@n$ in $u$}.
 | 
|  |    143 | \item[\texttt{case $e$ of $c@1$ => $e@1$ | \dots | $c@n$ => $e@n$}]
 | 
|  |    144 | \indexbold{*case}
 | 
|  |    145 | evaluates to $e@i$ if $e$ is of the form
 | 
|  |    146 | $c@i$. See~\S\ref{sec:case-expressions} for details.
 | 
|  |    147 | \end{description}
 | 
|  |    148 | 
 | 
|  |    149 | Terms may also contain $\lambda$-abstractions. For example, $\lambda
 | 
|  |    150 | x. x+1$ is the function that takes an argument $x$ and returns $x+1$. In
 | 
|  |    151 | Isabelle we write \texttt{\%x.~x+1}.\index{==>@{\tt\%}|bold}
 | 
|  |    152 | Instead of \texttt{\%x.~\%y.~\%z.~t} we can write \texttt{\%x~y~z.~t}.
 | 
|  |    153 | 
 | 
|  |    154 | \textbf{Formulae}\indexbold{formula}
 | 
|  |    155 | are terms of type \texttt{bool}. There are the basic
 | 
|  |    156 | constants \ttindexbold{True} and \ttindexbold{False} and the usual logical
 | 
|  |    157 | connectives (in decreasing order of priority):
 | 
|  |    158 | \verb$~$\index{$HOL1@{\ttnot}|bold} (`not'),
 | 
|  |    159 | \texttt{\&}\index{$HOL2@{\tt\&}|bold} (`and'),
 | 
|  |    160 | \texttt{|}\index{$HOL2@{\ttor}|bold} (`or') and
 | 
|  |    161 | \texttt{-->}\index{$HOL2@{\tt-->}|bold} (`implies'),
 | 
|  |    162 | all of which (except the unary \verb$~$) associate to the right. In
 | 
|  |    163 | particular \texttt{A --> B --> C} means \texttt{A --> (B --> C)} and is thus
 | 
|  |    164 | logically equivalent with \texttt{A \& B --> C}
 | 
|  |    165 | (which is \texttt{(A \& B) --> C}).
 | 
|  |    166 | 
 | 
|  |    167 | Equality is available in the form of the infix function
 | 
|  |    168 | \texttt{=} of type \texttt{'a => 'a => bool}. Thus \texttt{$t@1$ = $t@2$} is
 | 
|  |    169 | a formula provided $t@1$ and $t@2$ are terms of the same type. In case $t@1$
 | 
|  |    170 | and $t@2$ are of type \texttt{bool}, \texttt{=} acts as if-and-only-if.
 | 
|  |    171 | 
 | 
|  |    172 | The syntax for quantifiers is
 | 
|  |    173 | \texttt{!~$x$.$\,P$}\index{$HOLQ@{\ttall}|bold} (`for all $x$') and
 | 
|  |    174 | \texttt{?~$x$.$\,P$}\index{$HOLQ@{\tt?}|bold} (`exists $x$').
 | 
|  |    175 | There is even \texttt{?!~$x$.$\,P$}\index{$HOLQ@{\ttuniquex}|bold}, which
 | 
|  |    176 | means that there exists exactly one $x$ that satisfies $P$. Instead of
 | 
|  |    177 | \texttt{!} and \texttt{?} you may also write \texttt{ALL} and \texttt{EX}.
 | 
|  |    178 | Nested quantifications can be abbreviated:
 | 
|  |    179 | \texttt{!$x~y~z$.$\,P$} means \texttt{!$x$.~!$y$.~!$z$.$\,P$}.
 | 
|  |    180 | 
 | 
|  |    181 | Despite type inference, it is sometimes necessary to attach explicit
 | 
|  |    182 | \bfindex{type constraints} to a term.  The syntax is \texttt{$t$::$\tau$} as
 | 
|  |    183 | in \texttt{x < (y::nat)}. Note that \texttt{::} binds weakly and should
 | 
|  |    184 | therefore be enclosed in parentheses: \texttt{x < y::nat} is ill-typed
 | 
|  |    185 | because it is interpreted as \texttt{(x < y)::nat}. The main reason for type
 | 
|  |    186 | constraints are overloaded functions like \texttt{+}, \texttt{*} and
 | 
|  |    187 | \texttt{<}. (See \S\ref{sec:TypeClasses} for a full discussion of
 | 
|  |    188 | overloading.)
 | 
|  |    189 | 
 | 
|  |    190 | \begin{warn}
 | 
|  |    191 | In general, HOL's concrete syntax tries to follow the conventions of
 | 
|  |    192 | functional programming and mathematics. Below we list the main rules that you
 | 
|  |    193 | should be familiar with to avoid certain syntactic traps. A particular
 | 
|  |    194 | problem for novices can be the priority of operators. If you are unsure, use
 | 
|  |    195 | more rather than fewer parentheses. In those cases where Isabelle echoes your
 | 
|  |    196 | input, you can see which parentheses are dropped---they were superfluous. If
 | 
|  |    197 | you are unsure how to interpret Isabelle's output because you don't know
 | 
|  |    198 | where the (dropped) parentheses go, set (and possibly reset) the flag
 | 
|  |    199 | \ttindexbold{show_brackets}:
 | 
|  |    200 | \begin{ttbox}
 | 
|  |    201 | set show_brackets; \(\dots\); reset show_brackets;
 | 
|  |    202 | \end{ttbox}
 | 
|  |    203 | \end{warn}
 | 
|  |    204 | 
 | 
|  |    205 | \begin{itemize}
 | 
|  |    206 | \item
 | 
|  |    207 | Remember that \texttt{f t u} means \texttt{(f t) u} and not \texttt{f(t u)}!
 | 
|  |    208 | \item
 | 
|  |    209 | Isabelle allows infix functions like \texttt{+}. The prefix form of function
 | 
|  |    210 | application binds more strongly than anything else and hence \texttt{f~x + y}
 | 
|  |    211 | means \texttt{(f~x)~+~y} and not \texttt{f(x+y)}.
 | 
|  |    212 | \item
 | 
|  |    213 | Remember that in HOL if-and-only-if is expressed using equality.  But equality
 | 
|  |    214 | has a high priority, as befitting a  relation, while if-and-only-if typically
 | 
|  |    215 | has the lowest priority.  Thus, \verb$~ ~ P = P$ means \verb$~ ~(P = P)$ and
 | 
|  |    216 | not \verb$(~ ~P) = P$. When using \texttt{=} to mean logical equivalence,
 | 
|  |    217 | enclose both operands in parentheses, as in \texttt{(A \& B) = (B \& A)}.
 | 
|  |    218 | \item
 | 
|  |    219 | Constructs with an opening but without a closing delimiter bind very weakly
 | 
|  |    220 | and should therefore be enclosed in parentheses if they appear in subterms, as
 | 
|  |    221 | in \texttt{f = (\%x.~x)}. This includes
 | 
|  |    222 | \ttindex{if}, \ttindex{let}, \ttindex{case}, \verb$%$ and quantifiers.
 | 
|  |    223 | \item
 | 
|  |    224 | Never write \texttt{\%x.x} or \texttt{!x.x=x} because \texttt{x.x} is always
 | 
|  |    225 | read as a single qualified identifier that refers to an item \texttt{x} in
 | 
|  |    226 | theory \texttt{x}. Write \texttt{\%x.~x} and \texttt{!x.~x=x} instead.
 | 
|  |    227 | \end{itemize}
 | 
|  |    228 | 
 | 
|  |    229 | \section{Variables}
 | 
|  |    230 | \label{sec:variables}
 | 
|  |    231 | 
 | 
|  |    232 | Isabelle distinguishes free and bound variables just as is customary. Bound
 | 
|  |    233 | variables are automatically renamed to avoid clashes with free variables. In
 | 
|  |    234 | addition, Isabelle has a third kind of variable, called a \bfindex{schematic
 | 
|  |    235 |   variable} or \bfindex{unknown}, which starts with a \texttt{?}.  Logically,
 | 
|  |    236 | an unknown is a free variable. But it may be instantiated by another term
 | 
|  |    237 | during the proof process. For example, the mathematical theorem $x = x$ is
 | 
|  |    238 | represented in Isabelle as \texttt{?x = ?x}, which means that Isabelle can
 | 
|  |    239 | instantiate it arbitrarily. This is in contrast to ordinary variables, which
 | 
|  |    240 | remain fixed. The programming language Prolog calls unknowns {\em logical\/}
 | 
|  |    241 | variables.
 | 
|  |    242 | 
 | 
|  |    243 | Most of the time you can and should ignore unknowns and work with ordinary
 | 
|  |    244 | variables. Just don't be surprised that after you have finished the
 | 
| 6691 |    245 | proof of a theorem, Isabelle (i.e.\ \ttindex{qed} at the end of a proof) will
 | 
|  |    246 | turn your free variables into unknowns: it merely indicates that Isabelle will
 | 
|  |    247 | automatically instantiate those unknowns suitably when the theorem is used in
 | 
|  |    248 | some other proof.
 | 
| 5375 |    249 | \begin{warn}
 | 
|  |    250 |   The existential quantifier \texttt{?}\index{$HOLQ@{\tt?}} needs to be
 | 
|  |    251 |   followed by a space. Otherwise \texttt{?x} is interpreted as a schematic
 | 
|  |    252 |   variable.
 | 
|  |    253 | \end{warn}
 | 
|  |    254 | 
 | 
|  |    255 | \section{Getting started}
 | 
|  |    256 | 
 | 
|  |    257 | Assuming you have installed Isabelle, you start it by typing \texttt{isabelle
 | 
| 6584 |    258 |   HOL} in a shell window.\footnote{Simply executing \texttt{isabelle} without
 | 
|  |    259 |   an argument starts the default logic, which usually is already \texttt{HOL}.
 | 
|  |    260 |   This is controlled by the \texttt{ISABELLE_LOGIC} setting, see \emph{The
 | 
|  |    261 |     Isabelle System Manual} for more details.} This presents you with
 | 
|  |    262 | Isabelle's most basic ASCII interface.  In addition you need to open an editor
 | 
|  |    263 | window to create theories (\texttt{.thy} files) and proof scripts
 | 
|  |    264 | (\texttt{.ML} files). While you are developing a proof, we recommend to type
 | 
|  |    265 | each proof command into the ML-file first and then enter it into Isabelle by
 | 
|  |    266 | copy-and-paste, thus ensuring that you have a complete record of your proof.
 | 
|  |    267 | 
 |