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