CookBook/Package/Ind_Intro.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 29 Oct 2008 21:46:33 +0100
changeset 44 dee4b3e66dfe
parent 32 5bb2d29553c2
child 88 ebbd0dd008c8
permissions -rw-r--r--
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     1
theory Ind_Intro
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     2
imports Main
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     3
begin
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     4
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     5
chapter {* How to write a definitional package *}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     6
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     7
section{* Introduction *}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     8
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     9
text {*
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    10
\begin{flushright}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    11
  {\em
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    12
    ``My thesis is that programming is not at the bottom of the intellectual \\
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    13
    pyramid, but at the top. It's creative design of the highest order. It \\
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    14
    isn't monkey or donkey work; rather, as Edsger Dijkstra famously \\
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    15
    claimed, it's amongst the hardest intellectual tasks ever attempted.''} \\[1ex]
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    16
    Richard Bornat, In defence of programming
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    17
\end{flushright}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    18
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    19
\medskip
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    20
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    21
\noindent
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    22
Higher order logic, as implemented in Isabelle/HOL, is based on just a few primitive
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    23
constants, like equality, implication, and the description operator, whose properties are
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    24
described as axioms. All other concepts, such as inductive predicates, datatypes, or
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    25
recursive functions are \emph{defined} using these constants, and the desired
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    26
properties, for example induction theorems, or recursion equations are \emph{derived}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    27
from the definitions by a \emph{formal proof}. Since it would be very tedious
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    28
for the average user to define complex inductive predicates or datatypes ``by hand''
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    29
just using the primitive operators of higher order logic, Isabelle/HOL already contains
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    30
a number of \emph{packages} automating such tedious work. Thanks to those packages,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    31
the user can give a high-level specification, like a list of introduction rules or
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    32
constructors, and the package then does all the low-level definitions and proofs
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    33
behind the scenes. The packages are written in Standard ML, the implementation
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    34
language of Isabelle, and can be invoked by the user from within theory documents
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    35
written in the Isabelle/Isar language via specific commands. Most of the time,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    36
when using Isabelle for applications, users do not have to worry about the inner
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    37
workings of packages, since they can just use the packages that are already part
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    38
of the Isabelle distribution. However, when developing a general theory that is
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    39
intended to be applied by other users, one may need to write a new package from
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    40
scratch. Recent examples of such packages include the verification environment
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    41
for sequential imperative programs by Schirmer \cite{Schirmer-LPAR04}, the
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    42
package for defining general recursive functions by Krauss \cite{Krauss-IJCAR06},
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    43
as well as the nominal datatype package by Berghofer and Urban \cite{Urban-Berghofer06}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    44
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    45
The scientific value of implementing a package should not be underestimated:
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    46
it is often more than just the automation of long-established scientific
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    47
results. Of course, a carefully-developed theory on paper is indispensable
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    48
as a basis. However, without an implementation, such a theory will only be of
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    49
very limited practical use, since only an implementation enables other users
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    50
to apply the theory on a larger scale without too much effort. Moreover,
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    51
implementing a package is a bit like formalizing a paper proof in a theorem
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    52
prover. In the literature, there are many examples of paper proofs that
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    53
turned out to be incomplete or even faulty, and doing a formalization is
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    54
a good way of uncovering such errors and ensuring that a proof is really
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    55
correct. The same applies to the theory underlying definitional packages.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    56
For example, the general form of some complicated induction rules for nominal
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    57
datatypes turned out to be quite hard to get right on the first try, so
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    58
an implementation is an excellent way to find out whether the rules really
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    59
work in practice.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    60
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    61
Writing a package is a particularly difficult task for users that are new
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    62
to Isabelle, since its programming interface consists of thousands of functions.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    63
Rather than just listing all those functions, we give a step-by-step tutorial
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    64
for writing a package, using an example that is still simple enough to be
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    65
easily understandable, but at the same time sufficiently complex to demonstrate
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    66
enough of Isabelle's interesting features. As a running example, we have
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    67
chosen a rather simple package for defining inductive predicates. To keep
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    68
things simple, we will not use the general Knaster-Tarski fixpoint
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    69
theorem on complete lattices, which forms the basis of Isabelle's standard
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    70
inductive definition package originally due to Paulson \cite{Paulson-ind-defs}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    71
Instead, we will use a simpler \emph{impredicative} (i.e.\ involving
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    72
quantification on predicate variables) encoding of inductive predicates
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    73
suggested by Melham \cite{Melham:1992:PIR}. Due to its simplicity, this
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    74
package will necessarily have a reduced functionality. It does neither
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    75
support introduction rules involving arbitrary monotone operators, nor does
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    76
it prove case analysis (or inversion) rules. Moreover, it only proves a weaker
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    77
form of the rule induction theorem.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    78
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    79
Reading this article does not require any prior knowledge of Isabelle's programming
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    80
interface. However, we assume the reader to already be familiar with writing
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    81
proofs in Isabelle/HOL using the Isar language. For further information on
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    82
this topic, consult the book by Nipkow, Paulson, and Wenzel
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    83
\cite{isa-tutorial}. Moreover, in order to understand the pieces of
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    84
code given in this tutorial, some familiarity with the basic concepts of the Standard
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    85
ML programming language, as described for example in the textbook by Paulson
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    86
\cite{paulson-ml2}, is required as well.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    87
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    88
The rest of this article is structured as follows. In \S\ref{sec:ind-examples},
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    89
we will illustrate the ``manual'' definition of inductive predicates using
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    90
some examples. Starting from these examples, we will describe in
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    91
\S\ref{sec:ind-general-method} how the construction works in general.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    92
The following sections are then dedicated to the implementation of a
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    93
package that carries out the construction of such inductive predicates.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    94
First of all, a parser for a high-level notation for specifying inductive
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    95
predicates via a list of introduction rules is developed in \S\ref{sec:ind-interface}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    96
Having parsed the specification, a suitable primitive definition must be
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    97
added to the theory, which will be explained in \S\ref{sec:ind-definition}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    98
Finally, \S\ref{sec:ind-proofs} will focus on methods for proving introduction
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    99
and induction rules from the definitions introduced in \S\ref{sec:ind-definition}.
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   100
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   101
\nocite{Bornat-lecture}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   102
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   103
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   104
end