| 
13999
 | 
     1  | 
\section{Introduction}
 | 
| 
 | 
     2  | 
  | 
| 
25427
 | 
     3  | 
This is a tutorial introduction to structured proofs in Isabelle/HOL.
  | 
| 
 | 
     4  | 
It introduces the core of the proof language Isar by example. Isar is
  | 
| 
 | 
     5  | 
an extension of the \isa{apply}-style proofs introduced in the
 | 
| 
 | 
     6  | 
Isabelle/HOL tutorial~\cite{LNCS2283} with structured proofs in a
 | 
| 
 | 
     7  | 
stylised language of mathematics.  These proofs are readable for both
  | 
| 
 | 
     8  | 
human and machine.
  | 
| 
13999
 | 
     9  | 
  | 
| 
 | 
    10  | 
\subsection{A first glimpse of Isar}
 | 
| 
 | 
    11  | 
Below you find a simplified grammar for Isar proofs.
  | 
| 
 | 
    12  | 
Parentheses are used for grouping and $^?$ indicates an optional item:
  | 
| 
 | 
    13  | 
\begin{center}
 | 
| 
 | 
    14  | 
\begin{tabular}{lrl}
 | 
| 
 | 
    15  | 
\emph{proof} & ::= & \isakeyword{proof} \emph{method}$^?$
 | 
| 
 | 
    16  | 
                     \emph{statement}*
 | 
| 
 | 
    17  | 
                     \isakeyword{qed} \\
 | 
| 
 | 
    18  | 
                 &$\mid$& \isakeyword{by} \emph{method}\\[1ex]
 | 
| 
 | 
    19  | 
\emph{statement} &::= & \isakeyword{fix} \emph{variables} \\
 | 
| 
 | 
    20  | 
             &$\mid$& \isakeyword{assume} \emph{propositions} \\
 | 
| 
25427
 | 
    21  | 
             &$\mid$& (\isakeyword{from} \emph{fact}*)$^?$ 
 | 
| 
13999
 | 
    22  | 
                    (\isakeyword{show} $\mid$ \isakeyword{have})
 | 
| 
 | 
    23  | 
                      \emph{propositions} \emph{proof} \\[1ex]
 | 
| 
 | 
    24  | 
\emph{proposition} &::=& (\emph{label}{\bf:})$^?$ \emph{string} \\[1ex]
 | 
| 
 | 
    25  | 
\emph{fact} &::=& \emph{label}
 | 
| 
 | 
    26  | 
\end{tabular}
 | 
| 
 | 
    27  | 
\end{center}
 | 
| 
 | 
    28  | 
A proof can be either compound (\isakeyword{proof} --
 | 
| 
 | 
    29  | 
\isakeyword{qed}) or atomic (\isakeyword{by}). A \emph{method} is a
 | 
| 
25427
 | 
    30  | 
proof method.
  | 
| 
13999
 | 
    31  | 
  | 
| 
 | 
    32  | 
This is a typical proof skeleton:
  | 
| 
 | 
    33  | 
\begin{center}
 | 
| 
 | 
    34  | 
\begin{tabular}{@{}ll}
 | 
| 
 | 
    35  | 
\isakeyword{proof}\\
 | 
| 
 | 
    36  | 
\hspace*{3ex}\isakeyword{assume} \isa{"}\emph{the-assm}\isa{"}\\
 | 
| 
 | 
    37  | 
\hspace*{3ex}\isakeyword{have} \isa{"}\dots\isa{"} & --- intermediate result\\
 | 
| 
 | 
    38  | 
\hspace*{3ex}\vdots\\
 | 
| 
 | 
    39  | 
\hspace*{3ex}\isakeyword{have} \isa{"}\dots\isa{"} & --- intermediate result\\
 | 
| 
 | 
    40  | 
\hspace*{3ex}\isakeyword{show} \isa{"}\emph{the-concl}\isa{"}\\
 | 
| 
 | 
    41  | 
\isakeyword{qed}
 | 
| 
 | 
    42  | 
\end{tabular}
 | 
| 
 | 
    43  | 
\end{center}
 | 
| 
 | 
    44  | 
It proves \emph{the-assm}~$\Longrightarrow$~\emph{the-concl}. Text starting with
 | 
| 
 | 
    45  | 
``---'' is a comment. The intermediate \isakeyword{have}s are only
 | 
| 
 | 
    46  | 
there to bridge the gap between the assumption and the conclusion and
  | 
| 
 | 
    47  | 
do not contribute to the theorem being proved. In contrast,
  | 
| 
 | 
    48  | 
\isakeyword{show} establishes the conclusion of the theorem.
 | 
| 
 | 
    49  | 
  | 
| 
25427
 | 
    50  | 
\subsection{Background}
 | 
| 
 | 
    51  | 
  | 
| 
 | 
    52  | 
Interactive theorem proving has been dominated by a model of proof
  | 
| 
 | 
    53  | 
that goes back to the LCF system~\cite{LCF}: a proof is a more or less
 | 
| 
 | 
    54  | 
structured sequence of commands that manipulate an implicit proof
  | 
| 
 | 
    55  | 
state. Thus the proof text is only suitable for the machine; for a
  | 
| 
 | 
    56  | 
human, the proof only comes alive when he can see the state changes
  | 
| 
 | 
    57  | 
caused by the stepwise execution of the commands. Such proofs are like
  | 
| 
 | 
    58  | 
uncommented assembly language programs. Their Isabelle incarnation are
  | 
| 
 | 
    59  | 
sequences of \isa{apply}-commands.
 | 
| 
13999
 | 
    60  | 
  | 
| 
25427
 | 
    61  | 
In contrast there is the model of a mathematics-like proof language
  | 
| 
 | 
    62  | 
pioneered in the Mizar system~\cite{Rudnicki92} and followed by
 | 
| 
 | 
    63  | 
Isar~\cite{WenzelW-JAR}.
 | 
| 
 | 
    64  | 
The most important arguments in favour of this style are
  | 
| 
 | 
    65  | 
\emph{communication} and \emph{maintainance}: structured proofs are
 | 
| 
 | 
    66  | 
immensly more readable and maintainable than \isa{apply}-scripts.
 | 
| 
 | 
    67  | 
  | 
| 
 | 
    68  | 
For reading this tutorial you should be familiar with natural
  | 
| 
 | 
    69  | 
deduction and the basics of Isabelle/HOL~\cite{LNCS2283} although we
 | 
| 
 | 
    70  | 
summarize the most important aspects of Isabelle below.  The
  | 
| 
 | 
    71  | 
definitive Isar reference is its manual~\cite{Isar-Ref-Man}. For an
 | 
| 
 | 
    72  | 
example-based account of Isar's support for reasoning by chains of
  | 
| 
 | 
    73  | 
(in)equations see~\cite{BauerW-TPHOLs01}.
 | 
| 
 | 
    74  | 
  | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
\subsection{Bits of Isabelle}
 | 
| 
13999
 | 
    77  | 
  | 
| 
 | 
    78  | 
Isabelle's meta-logic comes with a type of \emph{propositions} with
 | 
| 
 | 
    79  | 
implication $\Longrightarrow$ and a universal quantifier $\bigwedge$ for expressing
  | 
| 
 | 
    80  | 
inference rules and generality.  Iterated implications $A_1 \Longrightarrow \dots
  | 
| 
 | 
    81  | 
A_n \Longrightarrow A$ may be abbreviated to $[\![ A_1; \dots; A_n ]\!] \Longrightarrow A$.
  | 
| 
 | 
    82  | 
Applying a theorem $A \Longrightarrow B$ (named \isa{T}) to a theorem $A$ (named
 | 
| 
 | 
    83  | 
\isa{U}) is written \isa{T[OF U]} and yields theorem $B$.
 | 
| 
 | 
    84  | 
  | 
| 
 | 
    85  | 
Isabelle terms are simply typed. Function types are
  | 
| 
 | 
    86  | 
written $\tau_1 \Rightarrow \tau_2$.
  | 
| 
 | 
    87  | 
  | 
| 
 | 
    88  | 
Free variables that may be instantiated (``logical variables'' in Prolog
  | 
| 
 | 
    89  | 
parlance) are prefixed with a \isa{?}. Typically, theorems are stated with
 | 
| 
 | 
    90  | 
ordinary free variables but after the proof those are automatically replaced
  | 
| 
 | 
    91  | 
by \isa{?}-variables. Thus the theorem can be used with arbitrary instances
 | 
| 
 | 
    92  | 
of its free variables.
  | 
| 
 | 
    93  | 
  | 
| 
 | 
    94  | 
Isabelle/HOL offers all the usual logical symbols like $\longrightarrow$, $\land$,
  | 
| 
 | 
    95  | 
$\forall$ etc. HOL formulae are propositions, e.g.\ $\forall$ can appear below
  | 
| 
 | 
    96  | 
$\Longrightarrow$, but not the other way around. Beware that $\longrightarrow$ binds more
  | 
| 
 | 
    97  | 
tightly than $\Longrightarrow$: in $\forall x. P \longrightarrow Q$ the $\forall x$ covers $P \longrightarrow Q$, whereas
  | 
| 
 | 
    98  | 
in $\forall x. P \Longrightarrow Q$ it covers only $P$.
  | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
Proof methods include \isa{rule} (which performs a backwards
 | 
| 
 | 
   101  | 
step with a given rule, unifying the conclusion of the rule with the
  | 
| 
 | 
   102  | 
current subgoal and replacing the subgoal by the premises of the
  | 
| 
 | 
   103  | 
rule), \isa{simp} (for simplification) and \isa{blast} (for predicate
 | 
| 
 | 
   104  | 
calculus reasoning).
  | 
| 
 | 
   105  | 
  | 
| 
25427
 | 
   106  | 
\subsection{Advice}
 | 
| 
13999
 | 
   107  | 
  | 
| 
25427
 | 
   108  | 
A word of warning for potential writers of Isar proofs.  It
  | 
| 
 | 
   109  | 
is easier to write obscure rather than readable texts.  Similarly,
  | 
| 
 | 
   110  | 
\isa{apply}-style proofs are sometimes easier to write than readable
 | 
| 
 | 
   111  | 
ones: structure does not emerge automatically but needs to be
  | 
| 
 | 
   112  | 
understood and imposed. If the precise structure of the proof is
  | 
| 
 | 
   113  | 
unclear at beginning, it can be useful to start with \isa{apply} for
 | 
| 
 | 
   114  | 
exploratory purposes until one has found a proof which can be
  | 
| 
 | 
   115  | 
converted into a structured text in a second step. Top down conversion
  | 
| 
 | 
   116  | 
is possible because Isar allows \isa{apply}-style proofs as components
 | 
| 
 | 
   117  | 
of structured ones.
  | 
| 
13999
 | 
   118  | 
  | 
| 
25427
 | 
   119  | 
Finally, do not be mislead by the simplicity of the formulae being proved,
  | 
| 
 | 
   120  | 
especially in the beginning. Isar has been used very successfully in
  | 
| 
 | 
   121  | 
large applications, for example the formalisation of Java
  | 
| 
 | 
   122  | 
dialects~\cite{KleinN-TOPLAS}.
 | 
| 
 | 
   123  | 
\medskip
  | 
| 
 | 
   124  | 
  | 
| 
 | 
   125  | 
The rest of this tutorial is divided into two parts.
  | 
| 
 | 
   126  | 
Section~\ref{sec:Logic} introduces proofs in pure logic based on
 | 
| 
 | 
   127  | 
natural deduction. Section~\ref{sec:Induct} is dedicated to induction.
 |