author | nipkow |
Mon, 06 Dec 1993 17:05:10 +0100 | |
changeset 189 | 831a9a7ab9f3 |
parent 181 | 9c2db771f224 |
child 355 | 77150178beb2 |
permissions | -rw-r--r-- |
181
9c2db771f224
removal of amssymbols.sty and lcp.sty; addition of iman.sty
lcp
parents:
179
diff
changeset
|
1 |
\documentstyle[12pt,a4,proof,iman,alltt]{article} |
179 | 2 |
\hyphenation{data-type} |
130 | 3 |
%CADE version should use 11pt and the Springer style |
103 | 4 |
\newif\ifCADE |
5 |
\CADEfalse |
|
6 |
||
130 | 7 |
\title{A Fixedpoint Approach to Implementing (Co)Inductive |
8 |
Definitions\thanks{Jim Grundy and Simon Thompson made detailed comments on |
|
9 |
a draft. Research funded by the SERC (grants GR/G53279, |
|
103 | 10 |
GR/H40570) and by the ESPRIT Basic Research Action 6453 `Types'.}} |
11 |
||
12 |
\author{{\em Lawrence C. Paulson}\\ |
|
13 |
Computer Laboratory, University of Cambridge} |
|
14 |
\date{\today} |
|
15 |
\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} |
|
16 |
||
17 |
\def\picture #1 by #2 (#3){% pictures: width by height (name) |
|
18 |
\message{Picture #3} |
|
19 |
\vbox to #2{\hrule width #1 height 0pt depth 0pt |
|
20 |
\vfill \special{picture #3}}} |
|
21 |
||
22 |
||
23 |
\newcommand\sbs{\subseteq} |
|
24 |
\let\To=\Rightarrow |
|
25 |
||
26 |
||
27 |
%%%\newcommand\Pow{{\tt Pow}} |
|
28 |
\let\pow=\wp |
|
29 |
\newcommand\RepFun{{\tt RepFun}} |
|
30 |
\newcommand\cons{{\tt cons}} |
|
31 |
\def\succ{{\tt succ}} |
|
32 |
\newcommand\split{{\tt split}} |
|
33 |
\newcommand\fst{{\tt fst}} |
|
34 |
\newcommand\snd{{\tt snd}} |
|
35 |
\newcommand\converse{{\tt converse}} |
|
36 |
\newcommand\domain{{\tt domain}} |
|
37 |
\newcommand\range{{\tt range}} |
|
38 |
\newcommand\field{{\tt field}} |
|
39 |
\newcommand\bndmono{\hbox{\tt bnd\_mono}} |
|
40 |
\newcommand\lfp{{\tt lfp}} |
|
41 |
\newcommand\gfp{{\tt gfp}} |
|
42 |
\newcommand\id{{\tt id}} |
|
43 |
\newcommand\trans{{\tt trans}} |
|
44 |
\newcommand\wf{{\tt wf}} |
|
45 |
\newcommand\wfrec{\hbox{\tt wfrec}} |
|
46 |
\newcommand\nat{{\tt nat}} |
|
47 |
\newcommand\natcase{\hbox{\tt nat\_case}} |
|
48 |
\newcommand\transrec{{\tt transrec}} |
|
49 |
\newcommand\rank{{\tt rank}} |
|
50 |
\newcommand\univ{{\tt univ}} |
|
51 |
\newcommand\Vrec{{\tt Vrec}} |
|
52 |
\newcommand\Inl{{\tt Inl}} |
|
53 |
\newcommand\Inr{{\tt Inr}} |
|
54 |
\newcommand\case{{\tt case}} |
|
55 |
\newcommand\lst{{\tt list}} |
|
56 |
\newcommand\Nil{{\tt Nil}} |
|
57 |
\newcommand\Cons{{\tt Cons}} |
|
58 |
\newcommand\lstcase{\hbox{\tt list\_case}} |
|
59 |
\newcommand\lstrec{\hbox{\tt list\_rec}} |
|
60 |
\newcommand\length{{\tt length}} |
|
61 |
\newcommand\listn{{\tt listn}} |
|
62 |
\newcommand\acc{{\tt acc}} |
|
63 |
\newcommand\primrec{{\tt primrec}} |
|
64 |
\newcommand\SC{{\tt SC}} |
|
65 |
\newcommand\CONST{{\tt CONST}} |
|
66 |
\newcommand\PROJ{{\tt PROJ}} |
|
67 |
\newcommand\COMP{{\tt COMP}} |
|
68 |
\newcommand\PREC{{\tt PREC}} |
|
69 |
||
70 |
\newcommand\quniv{{\tt quniv}} |
|
71 |
\newcommand\llist{{\tt llist}} |
|
72 |
\newcommand\LNil{{\tt LNil}} |
|
73 |
\newcommand\LCons{{\tt LCons}} |
|
74 |
\newcommand\lconst{{\tt lconst}} |
|
75 |
\newcommand\lleq{{\tt lleq}} |
|
76 |
\newcommand\map{{\tt map}} |
|
77 |
\newcommand\term{{\tt term}} |
|
78 |
\newcommand\Apply{{\tt Apply}} |
|
79 |
\newcommand\termcase{{\tt term\_case}} |
|
80 |
\newcommand\rev{{\tt rev}} |
|
81 |
\newcommand\reflect{{\tt reflect}} |
|
82 |
\newcommand\tree{{\tt tree}} |
|
83 |
\newcommand\forest{{\tt forest}} |
|
84 |
\newcommand\Part{{\tt Part}} |
|
85 |
\newcommand\TF{{\tt tree\_forest}} |
|
86 |
\newcommand\Tcons{{\tt Tcons}} |
|
87 |
\newcommand\Fcons{{\tt Fcons}} |
|
88 |
\newcommand\Fnil{{\tt Fnil}} |
|
89 |
\newcommand\TFcase{\hbox{\tt TF\_case}} |
|
90 |
\newcommand\Fin{{\tt Fin}} |
|
91 |
\newcommand\QInl{{\tt QInl}} |
|
92 |
\newcommand\QInr{{\tt QInr}} |
|
93 |
\newcommand\qsplit{{\tt qsplit}} |
|
94 |
\newcommand\qcase{{\tt qcase}} |
|
95 |
\newcommand\Con{{\tt Con}} |
|
96 |
\newcommand\data{{\tt data}} |
|
97 |
||
98 |
\binperiod %%%treat . like a binary operator |
|
99 |
||
100 |
\begin{document} |
|
101 |
\pagestyle{empty} |
|
102 |
\begin{titlepage} |
|
103 |
\maketitle |
|
104 |
\begin{abstract} |
|
105 |
Several theorem provers provide commands for formalizing recursive |
|
179 | 106 |
data\-types or inductively defined sets. This paper presents a new |
107 |
approach, based on fixedpoint definitions. It is unusually general: it |
|
108 |
admits all provably monotone inductive definitions. It is conceptually |
|
109 |
simple, which has allowed the easy implementation of mutual recursion and |
|
110 |
other conveniences. It also handles coinductive definitions: simply |
|
111 |
replace the least fixedpoint by a greatest fixedpoint. This represents |
|
112 |
the first automated support for coinductive definitions. |
|
130 | 113 |
|
114 |
The method has been implemented in Isabelle's formalization of ZF set |
|
179 | 115 |
theory. It should be applicable to any logic in which the Knaster-Tarski |
116 |
Theorem can be proved. The paper briefly describes a method of |
|
117 |
formalizing non-well-founded data structures in standard ZF set theory. |
|
103 | 118 |
|
119 |
Examples include lists of $n$ elements, the accessible part of a relation |
|
120 |
and the set of primitive recursive functions. One example of a |
|
130 | 121 |
coinductive definition is bisimulations for lazy lists. \ifCADE\else |
103 | 122 |
Recursive datatypes are examined in detail, as well as one example of a |
130 | 123 |
{\bf codatatype}: lazy lists. The appendices are simple user's manuals |
103 | 124 |
for this Isabelle/ZF package.\fi |
125 |
\end{abstract} |
|
126 |
% |
|
127 |
\begin{center} Copyright \copyright{} \number\year{} by Lawrence C. Paulson |
|
128 |
\end{center} |
|
129 |
\thispagestyle{empty} |
|
130 |
\end{titlepage} |
|
131 |
||
132 |
\tableofcontents |
|
133 |
\cleardoublepage |
|
130 | 134 |
\pagenumbering{arabic}\pagestyle{headings} |
103 | 135 |
|
136 |
\section{Introduction} |
|
137 |
Several theorem provers provide commands for formalizing recursive data |
|
138 |
structures, like lists and trees. Examples include Boyer and Moore's shell |
|
139 |
principle~\cite{bm79} and Melham's recursive type package for the HOL |
|
140 |
system~\cite{melham89}. Such data structures are called {\bf datatypes} |
|
141 |
below, by analogy with {\tt datatype} definitions in Standard~ML\@. |
|
142 |
||
130 | 143 |
A datatype is but one example of an {\bf inductive definition}. This |
103 | 144 |
specifies the least set closed under given rules~\cite{aczel77}. The |
145 |
collection of theorems in a logic is inductively defined. A structural |
|
146 |
operational semantics~\cite{hennessy90} is an inductive definition of a |
|
147 |
reduction or evaluation relation on programs. A few theorem provers |
|
148 |
provide commands for formalizing inductive definitions; these include |
|
149 |
Coq~\cite{paulin92} and again the HOL system~\cite{camilleri92}. |
|
150 |
||
130 | 151 |
The dual notion is that of a {\bf coinductive definition}. This specifies |
103 | 152 |
the greatest set closed under given rules. Important examples include |
153 |
using bisimulation relations to formalize equivalence of |
|
154 |
processes~\cite{milner89} or lazy functional programs~\cite{abramsky90}. |
|
155 |
Other examples include lazy lists and other infinite data structures; these |
|
130 | 156 |
are called {\bf codatatypes} below. |
103 | 157 |
|
158 |
Most existing implementations of datatype and inductive definitions accept |
|
130 | 159 |
an artificially narrow class of inputs, and are not easily extended. The |
103 | 160 |
shell principle and Coq's inductive definition rules are built into the |
161 |
underlying logic. Melham's packages derive datatypes and inductive |
|
162 |
definitions from specialized constructions in higher-order logic. |
|
163 |
||
164 |
This paper describes a package based on a fixedpoint approach. Least |
|
165 |
fixedpoints yield inductive definitions; greatest fixedpoints yield |
|
130 | 166 |
coinductive definitions. The package is uniquely powerful: |
103 | 167 |
\begin{itemize} |
168 |
\item It accepts the largest natural class of inductive definitions, namely |
|
130 | 169 |
all that are provably monotone. |
103 | 170 |
\item It accepts a wide class of datatype definitions. |
130 | 171 |
\item It handles coinductive and codatatype definitions. Most of |
172 |
the discussion below applies equally to inductive and coinductive |
|
103 | 173 |
definitions, and most of the code is shared. To my knowledge, this is |
130 | 174 |
the only package supporting coinductive definitions. |
103 | 175 |
\item Definitions may be mutually recursive. |
176 |
\end{itemize} |
|
177 |
The package is implemented in Isabelle~\cite{isabelle-intro}, using ZF set |
|
178 |
theory \cite{paulson-set-I,paulson-set-II}. However, the fixedpoint |
|
179 |
approach is independent of Isabelle. The recursion equations are specified |
|
180 |
as introduction rules for the mutually recursive sets. The package |
|
181 |
transforms these rules into a mapping over sets, and attempts to prove that |
|
182 |
the mapping is monotonic and well-typed. If successful, the package |
|
183 |
makes fixedpoint definitions and proves the introduction, elimination and |
|
130 | 184 |
(co)induction rules. The package consists of several Standard ML |
103 | 185 |
functors~\cite{paulson91}; it accepts its argument and returns its result |
186 |
as ML structures. |
|
187 |
||
188 |
Most datatype packages equip the new datatype with some means of expressing |
|
189 |
recursive functions. This is the main thing lacking from my package. Its |
|
190 |
fixedpoint operators define recursive sets, not recursive functions. But |
|
191 |
the Isabelle/ZF theory provides well-founded recursion and other logical |
|
192 |
tools to simplify this task~\cite{paulson-set-II}. |
|
193 |
||
130 | 194 |
{\bf Outline.} \S2 briefly introduces the least and greatest fixedpoint |
195 |
operators. \S3 discusses the form of introduction rules, mutual recursion and |
|
196 |
other points common to inductive and coinductive definitions. \S4 discusses |
|
197 |
induction and coinduction rules separately. \S5 presents several examples, |
|
198 |
including a coinductive definition. \S6 describes datatype definitions, while |
|
199 |
\S7 draws brief conclusions. \ifCADE\else The appendices are simple user's |
|
200 |
manuals for this Isabelle/ZF package.\fi |
|
103 | 201 |
|
202 |
Most of the definitions and theorems shown below have been generated by the |
|
203 |
package. I have renamed some variables to improve readability. |
|
204 |
||
205 |
\section{Fixedpoint operators} |
|
206 |
In set theory, the least and greatest fixedpoint operators are defined as |
|
207 |
follows: |
|
208 |
\begin{eqnarray*} |
|
209 |
\lfp(D,h) & \equiv & \inter\{X\sbs D. h(X)\sbs X\} \\ |
|
210 |
\gfp(D,h) & \equiv & \union\{X\sbs D. X\sbs h(X)\} |
|
211 |
\end{eqnarray*} |
|
130 | 212 |
Let $D$ be a set. Say that $h$ is {\bf bounded by}~$D$ if $h(D)\sbs D$, and |
213 |
{\bf monotone below~$D$} if |
|
103 | 214 |
$h(A)\sbs h(B)$ for all $A$ and $B$ such that $A\sbs B\sbs D$. If $h$ is |
215 |
bounded by~$D$ and monotone then both operators yield fixedpoints: |
|
216 |
\begin{eqnarray*} |
|
217 |
\lfp(D,h) & = & h(\lfp(D,h)) \\ |
|
218 |
\gfp(D,h) & = & h(\gfp(D,h)) |
|
219 |
\end{eqnarray*} |
|
220 |
These equations are instances of the Knaster-Tarski theorem, which states |
|
221 |
that every monotonic function over a complete lattice has a |
|
222 |
fixedpoint~\cite{davey&priestley}. It is obvious from their definitions |
|
223 |
that $\lfp$ must be the least fixedpoint, and $\gfp$ the greatest. |
|
224 |
||
225 |
This fixedpoint theory is simple. The Knaster-Tarski theorem is easy to |
|
226 |
prove. Showing monotonicity of~$h$ is trivial, in typical cases. We must |
|
227 |
also exhibit a bounding set~$D$ for~$h$. Sometimes this is trivial, as |
|
179 | 228 |
when a set of `theorems' is (co)inductively defined over some previously |
229 |
existing set of `formulae.' But defining the bounding set for |
|
130 | 230 |
(co)datatype definitions requires some effort; see~\S\ref{univ-sec} below. |
103 | 231 |
|
232 |
||
130 | 233 |
\section{Elements of an inductive or coinductive definition}\label{basic-sec} |
234 |
Consider a (co)inductive definition of the sets $R_1$, \ldots,~$R_n$, in |
|
103 | 235 |
mutual recursion. They will be constructed from previously existing sets |
236 |
$D_1$, \ldots,~$D_n$, respectively, which are called their {\bf domains}. |
|
237 |
The construction yields not $R_i\sbs D_i$ but $R_i\sbs D_1+\cdots+D_n$, where |
|
130 | 238 |
$R_i$ is contained in the image of~$D_i$ under an |
239 |
injection. Reasons for this are discussed |
|
240 |
elsewhere~\cite[\S4.5]{paulson-set-II}. |
|
103 | 241 |
|
242 |
The definition may involve arbitrary parameters $\vec{p}=p_1$, |
|
243 |
\ldots,~$p_k$. Each recursive set then has the form $R_i(\vec{p})$. The |
|
244 |
parameters must be identical every time they occur within a definition. This |
|
245 |
would appear to be a serious restriction compared with other systems such as |
|
246 |
Coq~\cite{paulin92}. For instance, we cannot define the lists of |
|
247 |
$n$ elements as the set $\listn(A,n)$ using rules where the parameter~$n$ |
|
130 | 248 |
varies. \S\ref{listn-sec} describes how to express this set using the |
249 |
inductive definition package. |
|
103 | 250 |
|
251 |
To avoid clutter below, the recursive sets are shown as simply $R_i$ |
|
252 |
instead of $R_i(\vec{p})$. |
|
253 |
||
254 |
\subsection{The form of the introduction rules}\label{intro-sec} |
|
255 |
The body of the definition consists of the desired introduction rules, |
|
256 |
specified as strings. The conclusion of each rule must have the form $t\in |
|
257 |
R_i$, where $t$ is any term. Premises typically have the same form, but |
|
258 |
they can have the more general form $t\in M(R_i)$ or express arbitrary |
|
259 |
side-conditions. |
|
260 |
||
261 |
The premise $t\in M(R_i)$ is permitted if $M$ is a monotonic operator on |
|
262 |
sets, satisfying the rule |
|
263 |
\[ \infer{M(A)\sbs M(B)}{A\sbs B} \] |
|
130 | 264 |
The user must supply the package with monotonicity rules for all such premises. |
103 | 265 |
|
266 |
Because any monotonic $M$ may appear in premises, the criteria for an |
|
267 |
acceptable definition is semantic rather than syntactic. A suitable choice |
|
268 |
of~$M$ and~$t$ can express a lot. The powerset operator $\pow$ is |
|
269 |
monotone, and the premise $t\in\pow(R)$ expresses $t\sbs R$; see |
|
130 | 270 |
\S\ref{acc-sec} for an example. The `list of' operator is monotone, as is |
271 |
easily proved by induction. The premise $t\in\lst(R)$ avoids having to |
|
272 |
encode the effect of~$\lst(R)$ using mutual recursion; see \S\ref{primrec-sec} |
|
273 |
and also my earlier paper~\cite[\S4.4]{paulson-set-II}. |
|
103 | 274 |
|
275 |
Introduction rules may also contain {\bf side-conditions}. These are |
|
276 |
premises consisting of arbitrary formulae not mentioning the recursive |
|
277 |
sets. Side-conditions typically involve type-checking. One example is the |
|
278 |
premise $a\in A$ in the following rule from the definition of lists: |
|
279 |
\[ \infer{\Cons(a,l)\in\lst(A)}{a\in A & l\in\lst(A)} \] |
|
280 |
||
281 |
\subsection{The fixedpoint definitions} |
|
282 |
The package translates the list of desired introduction rules into a fixedpoint |
|
283 |
definition. Consider, as a running example, the finite set operator |
|
284 |
$\Fin(A)$: the set of all finite subsets of~$A$. It can be |
|
285 |
defined as the least set closed under the rules |
|
286 |
\[ \emptyset\in\Fin(A) \qquad |
|
287 |
\infer{\{a\}\un b\in\Fin(A)}{a\in A & b\in\Fin(A)} |
|
288 |
\] |
|
289 |
||
130 | 290 |
The domain in a (co)inductive definition must be some existing set closed |
103 | 291 |
under the rules. A suitable domain for $\Fin(A)$ is $\pow(A)$, the set of all |
292 |
subsets of~$A$. The package generates the definition |
|
293 |
\begin{eqnarray*} |
|
294 |
\Fin(A) & \equiv & \lfp(\pow(A), \; |
|
295 |
\begin{array}[t]{r@{\,}l} |
|
296 |
\lambda X. \{z\in\pow(A). & z=\emptyset \disj{} \\ |
|
297 |
&(\exists a\,b. z=\{a\}\un b\conj a\in A\conj b\in X)\}) |
|
298 |
\end{array} |
|
130 | 299 |
\end{eqnarray*} |
103 | 300 |
The contribution of each rule to the definition of $\Fin(A)$ should be |
130 | 301 |
obvious. A coinductive definition is similar but uses $\gfp$ instead |
103 | 302 |
of~$\lfp$. |
303 |
||
304 |
The package must prove that the fixedpoint operator is applied to a |
|
305 |
monotonic function. If the introduction rules have the form described |
|
306 |
above, and if the package is supplied a monotonicity theorem for every |
|
307 |
$t\in M(R_i)$ premise, then this proof is trivial.\footnote{Due to the |
|
308 |
presence of logical connectives in the fixedpoint's body, the |
|
309 |
monotonicity proof requires some unusual rules. These state that the |
|
130 | 310 |
connectives $\conj$, $\disj$ and $\exists$ preserve monotonicity with respect |
311 |
to the partial ordering on unary predicates given by $P\sqsubseteq Q$ if and |
|
103 | 312 |
only if $\forall x.P(x)\imp Q(x)$.} |
313 |
||
130 | 314 |
The result structure contains the definitions of the recursive sets as a theorem |
315 |
list called {\tt defs}. It also contains, as the theorem {\tt unfold}, a |
|
103 | 316 |
fixedpoint equation such as |
317 |
\begin{eqnarray*} |
|
318 |
\Fin(A) & = & |
|
319 |
\begin{array}[t]{r@{\,}l} |
|
320 |
\{z\in\pow(A). & z=\emptyset \disj{} \\ |
|
321 |
&(\exists a\,b. z=\{a\}\un b\conj a\in A\conj b\in \Fin(A))\} |
|
322 |
\end{array} |
|
323 |
\end{eqnarray*} |
|
130 | 324 |
It also contains, as the theorem {\tt dom\_subset}, an inclusion such as |
103 | 325 |
$\Fin(A)\sbs\pow(A)$. |
326 |
||
327 |
||
328 |
\subsection{Mutual recursion} \label{mutual-sec} |
|
130 | 329 |
In a mutually recursive definition, the domain of the fixedpoint construction |
103 | 330 |
is the disjoint sum of the domain~$D_i$ of each~$R_i$, for $i=1$, |
331 |
\ldots,~$n$. The package uses the injections of the |
|
332 |
binary disjoint sum, typically $\Inl$ and~$\Inr$, to express injections |
|
130 | 333 |
$h_{1n}$, \ldots, $h_{nn}$ for the $n$-ary disjoint sum $D_1+\cdots+D_n$. |
103 | 334 |
|
335 |
As discussed elsewhere \cite[\S4.5]{paulson-set-II}, Isabelle/ZF defines the |
|
336 |
operator $\Part$ to support mutual recursion. The set $\Part(A,h)$ |
|
337 |
contains those elements of~$A$ having the form~$h(z)$: |
|
338 |
\begin{eqnarray*} |
|
339 |
\Part(A,h) & \equiv & \{x\in A. \exists z. x=h(z)\}. |
|
340 |
\end{eqnarray*} |
|
341 |
For mutually recursive sets $R_1$, \ldots,~$R_n$ with |
|
342 |
$n>1$, the package makes $n+1$ definitions. The first defines a set $R$ using |
|
343 |
a fixedpoint operator. The remaining $n$ definitions have the form |
|
344 |
\begin{eqnarray*} |
|
130 | 345 |
R_i & \equiv & \Part(R,h_{in}), \qquad i=1,\ldots, n. |
103 | 346 |
\end{eqnarray*} |
347 |
It follows that $R=R_1\un\cdots\un R_n$, where the $R_i$ are pairwise disjoint. |
|
348 |
||
349 |
||
350 |
\subsection{Proving the introduction rules} |
|
130 | 351 |
The user supplies the package with the desired form of the introduction |
103 | 352 |
rules. Once it has derived the theorem {\tt unfold}, it attempts |
130 | 353 |
to prove those rules. From the user's point of view, this is the |
103 | 354 |
trickiest stage; the proofs often fail. The task is to show that the domain |
355 |
$D_1+\cdots+D_n$ of the combined set $R_1\un\cdots\un R_n$ is |
|
356 |
closed under all the introduction rules. This essentially involves replacing |
|
357 |
each~$R_i$ by $D_1+\cdots+D_n$ in each of the introduction rules and |
|
358 |
attempting to prove the result. |
|
359 |
||
360 |
Consider the $\Fin(A)$ example. After substituting $\pow(A)$ for $\Fin(A)$ |
|
361 |
in the rules, the package must prove |
|
362 |
\[ \emptyset\in\pow(A) \qquad |
|
363 |
\infer{\{a\}\un b\in\pow(A)}{a\in A & b\in\pow(A)} |
|
364 |
\] |
|
365 |
Such proofs can be regarded as type-checking the definition. The user |
|
366 |
supplies the package with type-checking rules to apply. Usually these are |
|
367 |
general purpose rules from the ZF theory. They could however be rules |
|
368 |
specifically proved for a particular inductive definition; sometimes this is |
|
369 |
the easiest way to get the definition through! |
|
370 |
||
130 | 371 |
The result structure contains the introduction rules as the theorem list {\tt |
372 |
intrs}. |
|
103 | 373 |
|
374 |
\subsection{The elimination rule} |
|
130 | 375 |
The elimination rule, called {\tt elim}, performs case analysis: a |
376 |
case for each introduction rule. The elimination rule |
|
377 |
for $\Fin(A)$ is |
|
103 | 378 |
\[ \infer{Q}{x\in\Fin(A) & \infer*{Q}{[x=\emptyset]} |
379 |
& \infer*{Q}{[x=\{a\}\un b & a\in A &b\in\Fin(A)]_{a,b}} } |
|
380 |
\] |
|
130 | 381 |
This rule states that if $x\in\Fin(A)$ then either $x=\emptyset$ or else |
382 |
$x=\{a\}\un b$ for some $a\in A$ and $b\in\Fin(A)$; it is a simple consequence |
|
383 |
of {\tt unfold}. |
|
384 |
||
385 |
\ifCADE\typeout{****Omitting mk_cases from CADE version!}\else |
|
103 | 386 |
The package also returns a function {\tt mk\_cases}, for generating simplified |
387 |
instances of the elimination rule. However, {\tt mk\_cases} only works for |
|
388 |
datatypes and for inductive definitions involving datatypes, such as an |
|
389 |
inductively defined relation between lists. It instantiates {\tt elim} |
|
390 |
with a user-supplied term, then simplifies the cases using the freeness of |
|
391 |
the underlying datatype. |
|
130 | 392 |
See \S\ref{mkcases} for an example. |
393 |
\fi |
|
103 | 394 |
|
130 | 395 |
\section{Induction and coinduction rules} |
396 |
Here we must consider inductive and coinductive definitions separately. |
|
103 | 397 |
For an inductive definition, the package returns an induction rule derived |
398 |
directly from the properties of least fixedpoints, as well as a modified |
|
399 |
rule for mutual recursion and inductively defined relations. For a |
|
130 | 400 |
coinductive definition, the package returns a basic coinduction rule. |
103 | 401 |
|
402 |
\subsection{The basic induction rule}\label{basic-ind-sec} |
|
130 | 403 |
The basic rule, called {\tt induct}, is appropriate in most situations. |
103 | 404 |
For inductive definitions, it is strong rule induction~\cite{camilleri92}; for |
405 |
datatype definitions (see below), it is just structural induction. |
|
406 |
||
407 |
The induction rule for an inductively defined set~$R$ has the following form. |
|
408 |
The major premise is $x\in R$. There is a minor premise for each |
|
409 |
introduction rule: |
|
410 |
\begin{itemize} |
|
411 |
\item If the introduction rule concludes $t\in R_i$, then the minor premise |
|
412 |
is~$P(t)$. |
|
413 |
||
414 |
\item The minor premise's eigenvariables are precisely the introduction |
|
130 | 415 |
rule's free variables that are not parameters of~$R$. For instance, the |
416 |
eigenvariables in the $\Fin(A)$ rule below are $a$ and $b$, but not~$A$. |
|
103 | 417 |
|
418 |
\item If the introduction rule has a premise $t\in R_i$, then the minor |
|
419 |
premise discharges the assumption $t\in R_i$ and the induction |
|
420 |
hypothesis~$P(t)$. If the introduction rule has a premise $t\in M(R_i)$ |
|
421 |
then the minor premise discharges the single assumption |
|
422 |
\[ t\in M(\{z\in R_i. P(z)\}). \] |
|
423 |
Because $M$ is monotonic, this assumption implies $t\in M(R_i)$. The |
|
424 |
occurrence of $P$ gives the effect of an induction hypothesis, which may be |
|
425 |
exploited by appealing to properties of~$M$. |
|
426 |
\end{itemize} |
|
130 | 427 |
The induction rule for $\Fin(A)$ resembles the elimination rule shown above, |
428 |
but includes an induction hypothesis: |
|
103 | 429 |
\[ \infer{P(x)}{x\in\Fin(A) & P(\emptyset) |
430 |
& \infer*{P(\{a\}\un b)}{[a\in A & b\in\Fin(A) & P(b)]_{a,b}} } |
|
431 |
\] |
|
432 |
Stronger induction rules often suggest themselves. In the case of |
|
433 |
$\Fin(A)$, the Isabelle/ZF theory proceeds to derive a rule whose third |
|
434 |
premise discharges the extra assumption $a\not\in b$. Most special induction |
|
435 |
rules must be proved manually, but the package proves a rule for mutual |
|
436 |
induction and inductive relations. |
|
437 |
||
438 |
\subsection{Mutual induction} |
|
439 |
The mutual induction rule is called {\tt |
|
440 |
mutual\_induct}. It differs from the basic rule in several respects: |
|
441 |
\begin{itemize} |
|
442 |
\item Instead of a single predicate~$P$, it uses $n$ predicates $P_1$, |
|
443 |
\ldots,~$P_n$: one for each recursive set. |
|
444 |
||
445 |
\item There is no major premise such as $x\in R_i$. Instead, the conclusion |
|
446 |
refers to all the recursive sets: |
|
447 |
\[ (\forall z.z\in R_1\imp P_1(z))\conj\cdots\conj |
|
448 |
(\forall z.z\in R_n\imp P_n(z)) |
|
449 |
\] |
|
450 |
Proving the premises simultaneously establishes $P_i(z)$ for $z\in |
|
451 |
R_i$ and $i=1$, \ldots,~$n$. |
|
452 |
||
453 |
\item If the domain of some $R_i$ is the Cartesian product |
|
454 |
$A_1\times\cdots\times A_m$, then the corresponding predicate $P_i$ takes $m$ |
|
455 |
arguments and the corresponding conjunct of the conclusion is |
|
456 |
\[ (\forall z_1\ldots z_m.\pair{z_1,\ldots,z_m}\in R_i\imp P_i(z_1,\ldots,z_m)) |
|
457 |
\] |
|
458 |
\end{itemize} |
|
459 |
The last point above simplifies reasoning about inductively defined |
|
460 |
relations. It eliminates the need to express properties of $z_1$, |
|
461 |
\ldots,~$z_m$ as properties of the tuple $\pair{z_1,\ldots,z_m}$. |
|
462 |
||
130 | 463 |
\subsection{Coinduction}\label{coind-sec} |
464 |
A coinductive definition yields a primitive coinduction rule, with no |
|
103 | 465 |
refinements such as those for the induction rules. (Experience may suggest |
130 | 466 |
refinements later.) Consider the codatatype of lazy lists as an example. For |
103 | 467 |
suitable definitions of $\LNil$ and $\LCons$, lazy lists may be defined as the |
468 |
greatest fixedpoint satisfying the rules |
|
469 |
\[ \LNil\in\llist(A) \qquad |
|
470 |
\infer[(-)]{\LCons(a,l)\in\llist(A)}{a\in A & l\in\llist(A)} |
|
471 |
\] |
|
130 | 472 |
The $(-)$ tag stresses that this is a coinductive definition. A suitable |
103 | 473 |
domain for $\llist(A)$ is $\quniv(A)$, a set closed under variant forms of |
474 |
sum and product for representing infinite data structures |
|
130 | 475 |
(see~\S\ref{univ-sec}). Coinductive definitions use these variant sums and |
103 | 476 |
products. |
477 |
||
478 |
The package derives an {\tt unfold} theorem similar to that for $\Fin(A)$. |
|
479 |
Then it proves the theorem {\tt co\_induct}, which expresses that $\llist(A)$ |
|
480 |
is the greatest solution to this equation contained in $\quniv(A)$: |
|
130 | 481 |
\[ \infer{x\in\llist(A)}{x\in X & X\sbs \quniv(A) & |
103 | 482 |
\infer*{z=\LNil\disj \bigl(\exists a\,l.\, |
483 |
\begin{array}[t]{@{}l} |
|
484 |
z=\LCons(a,l) \conj a\in A \conj{}\\ |
|
485 |
l\in X\un\llist(A) \bigr) |
|
486 |
\end{array} }{[z\in X]_z}} |
|
487 |
\] |
|
130 | 488 |
This rule complements the introduction rules; it provides a means of showing |
489 |
$x\in\llist(A)$ when $x$ is infinite. For instance, if $x=\LCons(0,x)$ then |
|
490 |
applying the rule with $X=\{x\}$ proves $x\in\llist(\nat)$. Here $\nat$ |
|
491 |
is the set of natural numbers. |
|
492 |
||
103 | 493 |
Having $X\un\llist(A)$ instead of simply $X$ in the third premise above |
494 |
represents a slight strengthening of the greatest fixedpoint property. I |
|
130 | 495 |
discuss several forms of coinduction rules elsewhere~\cite{paulson-coind}. |
103 | 496 |
|
497 |
||
130 | 498 |
\section{Examples of inductive and coinductive definitions}\label{ind-eg-sec} |
103 | 499 |
This section presents several examples: the finite set operator, |
500 |
lists of $n$ elements, bisimulations on lazy lists, the well-founded part |
|
501 |
of a relation, and the primitive recursive functions. |
|
502 |
||
503 |
\subsection{The finite set operator} |
|
504 |
The definition of finite sets has been discussed extensively above. Here |
|
505 |
is the corresponding ML invocation (note that $\cons(a,b)$ abbreviates |
|
506 |
$\{a\}\un b$ in Isabelle/ZF): |
|
507 |
\begin{ttbox} |
|
508 |
structure Fin = Inductive_Fun |
|
509 |
(val thy = Arith.thy addconsts [(["Fin"],"i=>i")]; |
|
510 |
val rec_doms = [("Fin","Pow(A)")]; |
|
511 |
val sintrs = |
|
512 |
["0 : Fin(A)", |
|
513 |
"[| a: A; b: Fin(A) |] ==> cons(a,b) : Fin(A)"]; |
|
514 |
val monos = []; |
|
515 |
val con_defs = []; |
|
130 | 516 |
val type_intrs = [empty_subsetI, cons_subsetI, PowI]; |
103 | 517 |
val type_elims = [make_elim PowD]); |
518 |
\end{ttbox} |
|
519 |
The parent theory is obtained from {\tt Arith.thy} by adding the unary |
|
520 |
function symbol~$\Fin$. Its domain is specified as $\pow(A)$, where $A$ is |
|
521 |
the parameter appearing in the introduction rules. For type-checking, the |
|
522 |
package supplies the introduction rules: |
|
523 |
\[ \emptyset\sbs A \qquad |
|
524 |
\infer{\{a\}\un B\sbs C}{a\in C & B\sbs C} |
|
525 |
\] |
|
526 |
A further introduction rule and an elimination rule express the two |
|
527 |
directions of the equivalence $A\in\pow(B)\bimp A\sbs B$. Type-checking |
|
528 |
involves mostly introduction rules. When the package returns, we can refer |
|
529 |
to the $\Fin(A)$ introduction rules as {\tt Fin.intrs}, the induction rule |
|
530 |
as {\tt Fin.induct}, and so forth. |
|
531 |
||
532 |
\subsection{Lists of $n$ elements}\label{listn-sec} |
|
179 | 533 |
This has become a standard example of an inductive definition. Following |
534 |
Paulin-Mohring~\cite{paulin92}, we could attempt to define a new datatype |
|
535 |
$\listn(A,n)$, for lists of length~$n$, as an $n$-indexed family of sets. |
|
536 |
But her introduction rules |
|
103 | 537 |
\[ {\tt Niln}\in\listn(A,0) \qquad |
538 |
\infer{{\tt Consn}(n,a,l)\in\listn(A,\succ(n))} |
|
539 |
{n\in\nat & a\in A & l\in\listn(A,n)} |
|
540 |
\] |
|
541 |
are not acceptable to the inductive definition package: |
|
542 |
$\listn$ occurs with three different parameter lists in the definition. |
|
543 |
||
544 |
\begin{figure} |
|
545 |
\begin{small} |
|
546 |
\begin{verbatim} |
|
547 |
structure ListN = Inductive_Fun |
|
548 |
(val thy = ListFn.thy addconsts [(["listn"],"i=>i")]; |
|
549 |
val rec_doms = [("listn", "nat*list(A)")]; |
|
550 |
val sintrs = |
|
551 |
["<0,Nil> : listn(A)", |
|
130 | 552 |
"[| a: A; <n,l> : listn(A) |] ==> |
553 |
<succ(n), Cons(a,l)> : listn(A)"]; |
|
103 | 554 |
val monos = []; |
555 |
val con_defs = []; |
|
130 | 556 |
val type_intrs = nat_typechecks @ List.intrs @ [SigmaI]; |
103 | 557 |
val type_elims = [SigmaE2]); |
558 |
\end{verbatim} |
|
559 |
\end{small} |
|
560 |
\hrule |
|
561 |
\caption{Defining lists of $n$ elements} \label{listn-fig} |
|
562 |
\end{figure} |
|
563 |
||
564 |
There is an obvious way of handling this particular example, which may suggest |
|
565 |
a general approach to varying parameters. Here, we can take advantage of an |
|
566 |
existing datatype definition of $\lst(A)$, with constructors $\Nil$ |
|
567 |
and~$\Cons$. Then incorporate the number~$n$ into the inductive set itself, |
|
568 |
defining $\listn(A)$ as a relation. It consists of pairs $\pair{n,l}$ such |
|
569 |
that $n\in\nat$ and~$l\in\lst(A)$ and $l$ has length~$n$. In fact, |
|
570 |
$\listn(A)$ turns out to be the converse of the length function on~$\lst(A)$. |
|
571 |
The Isabelle/ZF introduction rules are |
|
572 |
\[ \pair{0,\Nil}\in\listn(A) \qquad |
|
573 |
\infer{\pair{\succ(n),\Cons(a,l)}\in\listn(A)} |
|
574 |
{a\in A & \pair{n,l}\in\listn(A)} |
|
575 |
\] |
|
576 |
Figure~\ref{listn-fig} presents the ML invocation. A theory of lists, |
|
577 |
extended with a declaration of $\listn$, is the parent theory. The domain |
|
578 |
is specified as $\nat\times\lst(A)$. The type-checking rules include those |
|
579 |
for 0, $\succ$, $\Nil$ and $\Cons$. Because $\listn(A)$ is a set of pairs, |
|
580 |
type-checking also requires introduction and elimination rules to express |
|
581 |
both directions of the equivalence $\pair{a,b}\in A\times B \bimp a\in A |
|
582 |
\conj b\in B$. |
|
583 |
||
584 |
The package returns introduction, elimination and induction rules for |
|
585 |
$\listn$. The basic induction rule, {\tt ListN.induct}, is |
|
586 |
\[ \infer{P(x)}{x\in\listn(A) & P(\pair{0,\Nil}) & |
|
587 |
\infer*{P(\pair{\succ(n),\Cons(a,l)})} |
|
588 |
{[a\in A & \pair{n,l}\in\listn(A) & P(\pair{n,l})]_{a,l,n}}} |
|
589 |
\] |
|
590 |
This rule requires the induction formula to be a |
|
591 |
unary property of pairs,~$P(\pair{n,l})$. The alternative rule, {\tt |
|
592 |
ListN.mutual\_induct}, uses a binary property instead: |
|
130 | 593 |
\[ \infer{\forall n\,l. \pair{n,l}\in\listn(A) \imp P(n,l)} |
103 | 594 |
{P(0,\Nil) & |
595 |
\infer*{P(\succ(n),\Cons(a,l))} |
|
596 |
{[a\in A & \pair{n,l}\in\listn(A) & P(n,l)]_{a,l,n}}} |
|
597 |
\] |
|
598 |
It is now a simple matter to prove theorems about $\listn(A)$, such as |
|
599 |
\[ \forall l\in\lst(A). \pair{\length(l),\, l}\in\listn(A) \] |
|
600 |
\[ \listn(A)``\{n\} = \{l\in\lst(A). \length(l)=n\} \] |
|
130 | 601 |
This latter result --- here $r``X$ denotes the image of $X$ under $r$ |
103 | 602 |
--- asserts that the inductive definition agrees with the obvious notion of |
603 |
$n$-element list. |
|
604 |
||
605 |
Unlike in Coq, the definition does not declare a new datatype. A `list of |
|
130 | 606 |
$n$ elements' really is a list and is subject to list operators such |
607 |
as append (concatenation). For example, a trivial induction on |
|
608 |
$\pair{m,l}\in\listn(A)$ yields |
|
103 | 609 |
\[ \infer{\pair{m\mathbin{+} m,\, l@l'}\in\listn(A)} |
610 |
{\pair{m,l}\in\listn(A) & \pair{m',l'}\in\listn(A)} |
|
611 |
\] |
|
612 |
where $+$ here denotes addition on the natural numbers and @ denotes append. |
|
613 |
||
614 |
\ifCADE\typeout{****Omitting mk_cases from CADE version!} |
|
615 |
\else |
|
130 | 616 |
\subsection{A demonstration of {\tt mk\_cases}}\label{mkcases} |
103 | 617 |
The elimination rule, {\tt ListN.elim}, is cumbersome: |
618 |
\[ \infer{Q}{x\in\listn(A) & |
|
619 |
\infer*{Q}{[x = \pair{0,\Nil}]} & |
|
620 |
\infer*{Q} |
|
621 |
{\left[\begin{array}{l} |
|
622 |
x = \pair{\succ(n),\Cons(a,l)} \\ |
|
623 |
a\in A \\ |
|
624 |
\pair{n,l}\in\listn(A) |
|
625 |
\end{array} \right]_{a,l,n}}} |
|
626 |
\] |
|
179 | 627 |
The ML function {\tt ListN.mk\_cases} generates simplified instances of |
628 |
this rule. It works by freeness reasoning on the list constructors: |
|
629 |
$\Cons(a,l)$ is injective in its two arguments and differs from~$\Nil$. If |
|
630 |
$x$ is $\pair{i,\Nil}$ or $\pair{i,\Cons(a,l)}$ then {\tt ListN.mk\_cases} |
|
631 |
deduces the corresponding form of~$i$. For example, |
|
103 | 632 |
\begin{ttbox} |
633 |
ListN.mk_cases List.con_defs "<i,Cons(a,l)> : listn(A)" |
|
634 |
\end{ttbox} |
|
130 | 635 |
yields a rule with only two premises: |
103 | 636 |
\[ \infer{Q}{\pair{i, \Cons(a,l)}\in\listn(A) & |
637 |
\infer*{Q} |
|
638 |
{\left[\begin{array}{l} |
|
639 |
i = \succ(n) \\ a\in A \\ \pair{n,l}\in\listn(A) |
|
640 |
\end{array} \right]_{n}}} |
|
641 |
\] |
|
642 |
The package also has built-in rules for freeness reasoning about $0$ |
|
643 |
and~$\succ$. So if $x$ is $\pair{0,l}$ or $\pair{\succ(i),l}$, then {\tt |
|
644 |
ListN.mk\_cases} can similarly deduce the corresponding form of~$l$. |
|
645 |
||
646 |
The function {\tt mk\_cases} is also useful with datatype definitions |
|
647 |
themselves. The version from the definition of lists, namely {\tt |
|
648 |
List.mk\_cases}, can prove the rule |
|
649 |
\[ \infer{Q}{\Cons(a,l)\in\lst(A) & |
|
650 |
& \infer*{Q}{[a\in A &l\in\lst(A)]} } |
|
651 |
\] |
|
652 |
The most important uses of {\tt mk\_cases} concern inductive definitions of |
|
130 | 653 |
evaluation relations. Then {\tt mk\_cases} supports case analysis on |
654 |
possible evaluations, for example to prove that evaluation is |
|
655 |
functional. |
|
103 | 656 |
\fi %CADE |
657 |
||
130 | 658 |
\subsection{A coinductive definition: bisimulations on lazy lists} |
659 |
This example anticipates the definition of the codatatype $\llist(A)$, which |
|
660 |
consists of finite and infinite lists over~$A$. Its constructors are $\LNil$ |
|
661 |
and |
|
662 |
$\LCons$, satisfying the introduction rules shown in~\S\ref{coind-sec}. |
|
103 | 663 |
Because $\llist(A)$ is defined as a greatest fixedpoint and uses the variant |
664 |
pairing and injection operators, it contains non-well-founded elements such as |
|
665 |
solutions to $\LCons(a,l)=l$. |
|
666 |
||
130 | 667 |
The next step in the development of lazy lists is to define a coinduction |
103 | 668 |
principle for proving equalities. This is done by showing that the equality |
669 |
relation on lazy lists is the greatest fixedpoint of some monotonic |
|
670 |
operation. The usual approach~\cite{pitts94} is to define some notion of |
|
671 |
bisimulation for lazy lists, define equivalence to be the greatest |
|
672 |
bisimulation, and finally to prove that two lazy lists are equivalent if and |
|
130 | 673 |
only if they are equal. The coinduction rule for equivalence then yields a |
674 |
coinduction principle for equalities. |
|
103 | 675 |
|
676 |
A binary relation $R$ on lazy lists is a {\bf bisimulation} provided $R\sbs |
|
677 |
R^+$, where $R^+$ is the relation |
|
130 | 678 |
\[ \{\pair{\LNil,\LNil}\} \un |
679 |
\{\pair{\LCons(a,l),\LCons(a,l')} . a\in A \conj \pair{l,l'}\in R\}. |
|
103 | 680 |
\] |
681 |
||
682 |
A pair of lazy lists are {\bf equivalent} if they belong to some bisimulation. |
|
130 | 683 |
Equivalence can be coinductively defined as the greatest fixedpoint for the |
103 | 684 |
introduction rules |
130 | 685 |
\[ \pair{\LNil,\LNil} \in\lleq(A) \qquad |
686 |
\infer[(-)]{\pair{\LCons(a,l),\LCons(a,l')} \in\lleq(A)} |
|
687 |
{a\in A & \pair{l,l'}\in \lleq(A)} |
|
103 | 688 |
\] |
130 | 689 |
To make this coinductive definition, we invoke \verb|CoInductive_Fun|: |
103 | 690 |
\begin{ttbox} |
130 | 691 |
structure LList_Eq = CoInductive_Fun |
103 | 692 |
(val thy = LList.thy addconsts [(["lleq"],"i=>i")]; |
130 | 693 |
val rec_doms = [("lleq", "llist(A) * llist(A)")]; |
103 | 694 |
val sintrs = |
130 | 695 |
["<LNil, LNil> : lleq(A)", |
696 |
"[| a:A; <l,l'>: lleq(A) |] ==> |
|
697 |
<LCons(a,l), LCons(a,l')>: lleq(A)"]; |
|
103 | 698 |
val monos = []; |
699 |
val con_defs = []; |
|
130 | 700 |
val type_intrs = LList.intrs @ [SigmaI]; |
701 |
val type_elims = [SigmaE2]); |
|
103 | 702 |
\end{ttbox} |
703 |
Again, {\tt addconsts} declares a constant for $\lleq$ in the parent theory. |
|
130 | 704 |
The domain of $\lleq(A)$ is $\llist(A)\times\llist(A)$. The type-checking |
705 |
rules include the introduction rules for lazy lists as well as rules |
|
706 |
for both directions of the equivalence |
|
707 |
$\pair{a,b}\in A\times B \bimp a\in A \conj b\in B$. |
|
103 | 708 |
|
709 |
The package returns the introduction rules and the elimination rule, as |
|
130 | 710 |
usual. But instead of induction rules, it returns a coinduction rule. |
103 | 711 |
The rule is too big to display in the usual notation; its conclusion is |
130 | 712 |
$x\in\lleq(A)$ and its premises are $x\in X$, |
713 |
${X\sbs\llist(A)\times\llist(A)}$ and |
|
714 |
\[ \infer*{z=\pair{\LNil,\LNil}\disj \bigl(\exists a\,l\,l'.\, |
|
103 | 715 |
\begin{array}[t]{@{}l} |
130 | 716 |
z=\pair{\LCons(a,l),\LCons(a,l')} \conj a\in A \conj{}\\ |
717 |
\pair{l,l'}\in X\un\lleq(A) \bigr) |
|
103 | 718 |
\end{array} }{[z\in X]_z} |
719 |
\] |
|
130 | 720 |
Thus if $x\in X$, where $X$ is a bisimulation contained in the |
721 |
domain of $\lleq(A)$, then $x\in\lleq(A)$. It is easy to show that |
|
103 | 722 |
$\lleq(A)$ is reflexive: the equality relation is a bisimulation. And |
723 |
$\lleq(A)$ is symmetric: its converse is a bisimulation. But showing that |
|
130 | 724 |
$\lleq(A)$ coincides with the equality relation takes some work. |
103 | 725 |
|
726 |
\subsection{The accessible part of a relation}\label{acc-sec} |
|
727 |
Let $\prec$ be a binary relation on~$D$; in short, $(\prec)\sbs D\times D$. |
|
728 |
The {\bf accessible} or {\bf well-founded} part of~$\prec$, written |
|
729 |
$\acc(\prec)$, is essentially that subset of~$D$ for which $\prec$ admits |
|
730 |
no infinite decreasing chains~\cite{aczel77}. Formally, $\acc(\prec)$ is |
|
731 |
inductively defined to be the least set that contains $a$ if it contains |
|
732 |
all $\prec$-predecessors of~$a$, for $a\in D$. Thus we need an |
|
733 |
introduction rule of the form |
|
734 |
\[ \infer{a\in\acc(\prec)}{\forall y.y\prec a\imp y\in\acc(\prec)} \] |
|
735 |
Paulin-Mohring treats this example in Coq~\cite{paulin92}, but it causes |
|
736 |
difficulties for other systems. Its premise does not conform to |
|
737 |
the structure of introduction rules for HOL's inductive definition |
|
738 |
package~\cite{camilleri92}. It is also unacceptable to Isabelle package |
|
130 | 739 |
(\S\ref{intro-sec}), but fortunately can be transformed into the acceptable |
103 | 740 |
form $t\in M(R)$. |
741 |
||
742 |
The powerset operator is monotonic, and $t\in\pow(R)$ is equivalent to |
|
743 |
$t\sbs R$. This in turn is equivalent to $\forall y\in t. y\in R$. To |
|
744 |
express $\forall y.y\prec a\imp y\in\acc(\prec)$ we need only find a |
|
745 |
term~$t$ such that $y\in t$ if and only if $y\prec a$. A suitable $t$ is |
|
746 |
the inverse image of~$\{a\}$ under~$\prec$. |
|
747 |
||
748 |
The ML invocation below follows this approach. Here $r$ is~$\prec$ and |
|
130 | 749 |
$\field(r)$ refers to~$D$, the domain of $\acc(r)$. (The field of a |
750 |
relation is the union of its domain and range.) Finally |
|
751 |
$r^{-}``\{a\}$ denotes the inverse image of~$\{a\}$ under~$r$. The package is |
|
752 |
supplied the theorem {\tt Pow\_mono}, which asserts that $\pow$ is monotonic. |
|
103 | 753 |
\begin{ttbox} |
754 |
structure Acc = Inductive_Fun |
|
755 |
(val thy = WF.thy addconsts [(["acc"],"i=>i")]; |
|
756 |
val rec_doms = [("acc", "field(r)")]; |
|
757 |
val sintrs = |
|
130 | 758 |
["[| r-``\{a\}: Pow(acc(r)); a: field(r) |] ==> a: acc(r)"]; |
103 | 759 |
val monos = [Pow_mono]; |
760 |
val con_defs = []; |
|
761 |
val type_intrs = []; |
|
762 |
val type_elims = []); |
|
763 |
\end{ttbox} |
|
764 |
The Isabelle theory proceeds to prove facts about $\acc(\prec)$. For |
|
765 |
instance, $\prec$ is well-founded if and only if its field is contained in |
|
766 |
$\acc(\prec)$. |
|
767 |
||
768 |
As mentioned in~\S\ref{basic-ind-sec}, a premise of the form $t\in M(R)$ |
|
769 |
gives rise to an unusual induction hypothesis. Let us examine the |
|
770 |
induction rule, {\tt Acc.induct}: |
|
771 |
\[ \infer{P(x)}{x\in\acc(r) & |
|
772 |
\infer*{P(a)}{[r^{-}``\{a\}\in\pow(\{z\in\acc(r).P(z)\}) & |
|
773 |
a\in\field(r)]_a}} |
|
774 |
\] |
|
775 |
The strange induction hypothesis is equivalent to |
|
776 |
$\forall y. \pair{y,a}\in r\imp y\in\acc(r)\conj P(y)$. |
|
777 |
Therefore the rule expresses well-founded induction on the accessible part |
|
778 |
of~$\prec$. |
|
779 |
||
780 |
The use of inverse image is not essential. The Isabelle package can accept |
|
781 |
introduction rules with arbitrary premises of the form $\forall |
|
782 |
\vec{y}.P(\vec{y})\imp f(\vec{y})\in R$. The premise can be expressed |
|
783 |
equivalently as |
|
130 | 784 |
\[ \{z\in D. P(\vec{y}) \conj z=f(\vec{y})\} \in \pow(R) \] |
103 | 785 |
provided $f(\vec{y})\in D$ for all $\vec{y}$ such that~$P(\vec{y})$. The |
786 |
following section demonstrates another use of the premise $t\in M(R)$, |
|
787 |
where $M=\lst$. |
|
788 |
||
789 |
\subsection{The primitive recursive functions}\label{primrec-sec} |
|
790 |
The primitive recursive functions are traditionally defined inductively, as |
|
791 |
a subset of the functions over the natural numbers. One difficulty is that |
|
792 |
functions of all arities are taken together, but this is easily |
|
793 |
circumvented by regarding them as functions on lists. Another difficulty, |
|
794 |
the notion of composition, is less easily circumvented. |
|
795 |
||
796 |
Here is a more precise definition. Letting $\vec{x}$ abbreviate |
|
797 |
$x_0,\ldots,x_{n-1}$, we can write lists such as $[\vec{x}]$, |
|
798 |
$[y+1,\vec{x}]$, etc. A function is {\bf primitive recursive} if it |
|
799 |
belongs to the least set of functions in $\lst(\nat)\to\nat$ containing |
|
800 |
\begin{itemize} |
|
801 |
\item The {\bf successor} function $\SC$, such that $\SC[y,\vec{x}]=y+1$. |
|
802 |
\item All {\bf constant} functions $\CONST(k)$, such that |
|
803 |
$\CONST(k)[\vec{x}]=k$. |
|
804 |
\item All {\bf projection} functions $\PROJ(i)$, such that |
|
805 |
$\PROJ(i)[\vec{x}]=x_i$ if $0\leq i<n$. |
|
806 |
\item All {\bf compositions} $\COMP(g,[f_0,\ldots,f_{m-1}])$, |
|
807 |
where $g$ and $f_0$, \ldots, $f_{m-1}$ are primitive recursive, |
|
808 |
such that |
|
809 |
\begin{eqnarray*} |
|
810 |
\COMP(g,[f_0,\ldots,f_{m-1}])[\vec{x}] & = & |
|
811 |
g[f_0[\vec{x}],\ldots,f_{m-1}[\vec{x}]]. |
|
812 |
\end{eqnarray*} |
|
813 |
||
814 |
\item All {\bf recursions} $\PREC(f,g)$, where $f$ and $g$ are primitive |
|
815 |
recursive, such that |
|
816 |
\begin{eqnarray*} |
|
817 |
\PREC(f,g)[0,\vec{x}] & = & f[\vec{x}] \\ |
|
818 |
\PREC(f,g)[y+1,\vec{x}] & = & g[\PREC(f,g)[y,\vec{x}],\, y,\, \vec{x}]. |
|
819 |
\end{eqnarray*} |
|
820 |
\end{itemize} |
|
821 |
Composition is awkward because it combines not two functions, as is usual, |
|
822 |
but $m+1$ functions. In her proof that Ackermann's function is not |
|
823 |
primitive recursive, Nora Szasz was unable to formalize this definition |
|
824 |
directly~\cite{szasz93}. So she generalized primitive recursion to |
|
825 |
tuple-valued functions. This modified the inductive definition such that |
|
826 |
each operation on primitive recursive functions combined just two functions. |
|
827 |
||
828 |
\begin{figure} |
|
130 | 829 |
\begin{small}\begin{verbatim} |
103 | 830 |
structure Primrec = Inductive_Fun |
831 |
(val thy = Primrec0.thy; |
|
832 |
val rec_doms = [("primrec", "list(nat)->nat")]; |
|
130 | 833 |
val ext = None; |
103 | 834 |
val sintrs = |
835 |
["SC : primrec", |
|
836 |
"k: nat ==> CONST(k) : primrec", |
|
837 |
"i: nat ==> PROJ(i) : primrec", |
|
130 | 838 |
"[| g: primrec; fs: list(primrec) |] ==> COMP(g,fs): primrec", |
839 |
"[| f: primrec; g: primrec |] ==> PREC(f,g): primrec"]; |
|
103 | 840 |
val monos = [list_mono]; |
841 |
val con_defs = [SC_def,CONST_def,PROJ_def,COMP_def,PREC_def]; |
|
130 | 842 |
val type_intrs = pr0_typechecks; |
103 | 843 |
val type_elims = []); |
130 | 844 |
\end{verbatim}\end{small} |
103 | 845 |
\hrule |
846 |
\caption{Inductive definition of the primitive recursive functions} |
|
847 |
\label{primrec-fig} |
|
848 |
\end{figure} |
|
849 |
\def\fs{{\it fs}} |
|
850 |
Szasz was using ALF, but Coq and HOL would also have problems accepting |
|
851 |
this definition. Isabelle's package accepts it easily since |
|
852 |
$[f_0,\ldots,f_{m-1}]$ is a list of primitive recursive functions and |
|
853 |
$\lst$ is monotonic. There are five introduction rules, one for each of |
|
854 |
the five forms of primitive recursive function. Note the one for $\COMP$: |
|
855 |
\[ \infer{\COMP(g,\fs)\in\primrec}{g\in\primrec & \fs\in\lst(\primrec)} \] |
|
856 |
The induction rule for $\primrec$ has one case for each introduction rule. |
|
857 |
Due to the use of $\lst$ as a monotone operator, the composition case has |
|
858 |
an unusual induction hypothesis: |
|
859 |
\[ \infer*{P(\COMP(g,\fs))} |
|
130 | 860 |
{[g\in\primrec & \fs\in\lst(\{z\in\primrec.P(z)\})]_{\fs,g}} \] |
103 | 861 |
The hypothesis states that $\fs$ is a list of primitive recursive functions |
862 |
satisfying the induction formula. Proving the $\COMP$ case typically requires |
|
863 |
structural induction on lists, yielding two subcases: either $\fs=\Nil$ or |
|
864 |
else $\fs=\Cons(f,\fs')$, where $f\in\primrec$, $P(f)$, and $\fs'$ is |
|
865 |
another list of primitive recursive functions satisfying~$P$. |
|
866 |
||
867 |
Figure~\ref{primrec-fig} presents the ML invocation. Theory {\tt |
|
868 |
Primrec0.thy} defines the constants $\SC$, etc.; their definitions |
|
869 |
consist of routine list programming and are omitted. The Isabelle theory |
|
870 |
goes on to formalize Ackermann's function and prove that it is not |
|
871 |
primitive recursive, using the induction rule {\tt Primrec.induct}. The |
|
872 |
proof follows Szasz's excellent account. |
|
873 |
||
874 |
ALF and Coq treat inductive definitions as datatypes, with a new |
|
875 |
constructor for each introduction rule. This forced Szasz to define a |
|
876 |
small programming language for the primitive recursive functions, and then |
|
877 |
define their semantics. But the Isabelle/ZF formulation defines the |
|
878 |
primitive recursive functions directly as a subset of the function set |
|
879 |
$\lst(\nat)\to\nat$. This saves a step and conforms better to mathematical |
|
880 |
tradition. |
|
881 |
||
882 |
||
130 | 883 |
\section{Datatypes and codatatypes}\label{data-sec} |
884 |
A (co)datatype definition is a (co)inductive definition with automatically |
|
885 |
defined constructors and a case analysis operator. The package proves that the |
|
103 | 886 |
case operator inverts the constructors, and can also prove freeness theorems |
887 |
involving any pair of constructors. |
|
888 |
||
889 |
||
130 | 890 |
\subsection{Constructors and their domain}\label{univ-sec} |
891 |
Conceptually, our two forms of definition are distinct: a (co)inductive |
|
892 |
definition selects a subset of an existing set, but a (co)datatype |
|
103 | 893 |
definition creates a new set. But the package reduces the latter to the |
894 |
former. A set having strong closure properties must serve as the domain |
|
130 | 895 |
of the (co)inductive definition. Constructing this set requires some |
896 |
theoretical effort. Consider first datatypes and then codatatypes. |
|
103 | 897 |
|
898 |
Isabelle/ZF defines the standard notion of Cartesian product $A\times B$, |
|
899 |
containing ordered pairs $\pair{a,b}$. Now the $m$-tuple |
|
900 |
$\pair{x_1,\ldots\,x_m}$ is the empty set~$\emptyset$ if $m=0$, simply |
|
901 |
$x_1$ if $m=1$, and $\pair{x_1,\pair{x_2,\ldots\,x_m}}$ if $m\geq2$. |
|
902 |
Isabelle/ZF also defines the disjoint sum $A+B$, containing injections |
|
903 |
$\Inl(a)\equiv\pair{0,a}$ and $\Inr(b)\equiv\pair{1,b}$. |
|
904 |
||
905 |
A datatype constructor $\Con(x_1,\ldots\,x_m)$ is defined to be |
|
906 |
$h(\pair{x_1,\ldots\,x_m})$, where $h$ is composed of $\Inl$ and~$\Inr$. |
|
907 |
In a mutually recursive definition, all constructors for the set~$R_i$ have |
|
130 | 908 |
the outer form~$h_{in}$, where $h_{in}$ is the injection described |
103 | 909 |
in~\S\ref{mutual-sec}. Further nested injections ensure that the |
910 |
constructors for~$R_i$ are pairwise distinct. |
|
911 |
||
912 |
Isabelle/ZF defines the set $\univ(A)$, which contains~$A$ and |
|
913 |
furthermore contains $\pair{a,b}$, $\Inl(a)$ and $\Inr(b)$ for $a$, |
|
914 |
$b\in\univ(A)$. In a typical datatype definition with set parameters |
|
915 |
$A_1$, \ldots, $A_k$, a suitable domain for all the recursive sets is |
|
916 |
$\univ(A_1\un\cdots\un A_k)$. This solves the problem for |
|
917 |
datatypes~\cite[\S4.2]{paulson-set-II}. |
|
918 |
||
919 |
The standard pairs and injections can only yield well-founded |
|
920 |
constructions. This eases the (manual!) definition of recursive functions |
|
130 | 921 |
over datatypes. But they are unsuitable for codatatypes, which typically |
103 | 922 |
contain non-well-founded objects. |
923 |
||
130 | 924 |
To support codatatypes, Isabelle/ZF defines a variant notion of ordered |
103 | 925 |
pair, written~$\pair{a;b}$. It also defines the corresponding variant |
926 |
notion of Cartesian product $A\otimes B$, variant injections $\QInl(a)$ |
|
927 |
and~$\QInr(b)$, and variant disjoint sum $A\oplus B$. Finally it defines |
|
928 |
the set $\quniv(A)$, which contains~$A$ and furthermore contains |
|
929 |
$\pair{a;b}$, $\QInl(a)$ and $\QInr(b)$ for $a$, $b\in\quniv(A)$. In a |
|
130 | 930 |
typical codatatype definition with set parameters $A_1$, \ldots, $A_k$, a |
931 |
suitable domain is $\quniv(A_1\un\cdots\un A_k)$. This approach using |
|
932 |
standard ZF set theory\footnote{No reference is available. Variant pairs |
|
933 |
are defined by $\pair{a;b}\equiv (a+b) \equiv (\{0\}\times a) \un |
|
934 |
(\{1\}\times b)$, where $\times$ is the Cartesian product for standard |
|
935 |
ordered pairs. Now |
|
103 | 936 |
$\pair{a;b}$ is injective and monotonic in its two arguments. |
937 |
Non-well-founded constructions, such as infinite lists, are constructed |
|
938 |
as least fixedpoints; the bounding set typically has the form |
|
939 |
$\univ(a_1\un\cdots\un a_k)$, where $a_1$, \ldots, $a_k$ are specified |
|
940 |
elements of the construction.} |
|
130 | 941 |
is an alternative to adopting Aczel's Anti-Foundation |
942 |
Axiom~\cite{aczel88}. |
|
103 | 943 |
|
944 |
\subsection{The case analysis operator} |
|
130 | 945 |
The (co)datatype package automatically defines a case analysis operator, |
179 | 946 |
called {\tt$R$\_case}. A mutually recursive definition still has only one |
947 |
operator, whose name combines those of the recursive sets: it is called |
|
948 |
{\tt$R_1$\_\ldots\_$R_n$\_case}. The case operator is analogous to those |
|
949 |
for products and sums. |
|
103 | 950 |
|
951 |
Datatype definitions employ standard products and sums, whose operators are |
|
952 |
$\split$ and $\case$ and satisfy the equations |
|
953 |
\begin{eqnarray*} |
|
954 |
\split(f,\pair{x,y}) & = & f(x,y) \\ |
|
955 |
\case(f,g,\Inl(x)) & = & f(x) \\ |
|
956 |
\case(f,g,\Inr(y)) & = & g(y) |
|
957 |
\end{eqnarray*} |
|
958 |
Suppose the datatype has $k$ constructors $\Con_1$, \ldots,~$\Con_k$. Then |
|
959 |
its case operator takes $k+1$ arguments and satisfies an equation for each |
|
960 |
constructor: |
|
961 |
\begin{eqnarray*} |
|
962 |
R\hbox{\_case}(f_1,\ldots,f_k, {\tt Con}_i(\vec{x})) & = & f_i(\vec{x}), |
|
963 |
\qquad i = 1, \ldots, k |
|
964 |
\end{eqnarray*} |
|
130 | 965 |
The case operator's definition takes advantage of Isabelle's representation |
966 |
of syntax in the typed $\lambda$-calculus; it could readily be adapted to a |
|
967 |
theorem prover for higher-order logic. If $f$ and~$g$ have meta-type |
|
968 |
$i\To i$ then so do $\split(f)$ and |
|
969 |
$\case(f,g)$. This works because $\split$ and $\case$ operate on their last |
|
970 |
argument. They are easily combined to make complex case analysis |
|
103 | 971 |
operators. Here are two examples: |
972 |
\begin{itemize} |
|
973 |
\item $\split(\lambda x.\split(f(x)))$ performs case analysis for |
|
974 |
$A\times (B\times C)$, as is easily verified: |
|
975 |
\begin{eqnarray*} |
|
976 |
\split(\lambda x.\split(f(x)), \pair{a,b,c}) |
|
977 |
& = & (\lambda x.\split(f(x))(a,\pair{b,c}) \\ |
|
978 |
& = & \split(f(a), \pair{b,c}) \\ |
|
979 |
& = & f(a,b,c) |
|
980 |
\end{eqnarray*} |
|
981 |
||
982 |
\item $\case(f,\case(g,h))$ performs case analysis for $A+(B+C)$; let us |
|
983 |
verify one of the three equations: |
|
984 |
\begin{eqnarray*} |
|
985 |
\case(f,\case(g,h), \Inr(\Inl(b))) |
|
986 |
& = & \case(g,h,\Inl(b)) \\ |
|
987 |
& = & g(b) |
|
988 |
\end{eqnarray*} |
|
989 |
\end{itemize} |
|
130 | 990 |
Codatatype definitions are treated in precisely the same way. They express |
103 | 991 |
case operators using those for the variant products and sums, namely |
992 |
$\qsplit$ and~$\qcase$. |
|
993 |
||
994 |
||
995 |
\ifCADE The package has processed all the datatypes discussed in my earlier |
|
130 | 996 |
paper~\cite{paulson-set-II} and the codatatype of lazy lists. Space |
103 | 997 |
limitations preclude discussing these examples here, but they are |
998 |
distributed with Isabelle. |
|
999 |
\typeout{****Omitting datatype examples from CADE version!} \else |
|
1000 |
||
1001 |
To see how constructors and the case analysis operator are defined, let us |
|
1002 |
examine some examples. These include lists and trees/forests, which I have |
|
1003 |
discussed extensively in another paper~\cite{paulson-set-II}. |
|
1004 |
||
1005 |
\begin{figure} |
|
1006 |
\begin{ttbox} |
|
1007 |
structure List = Datatype_Fun |
|
1008 |
(val thy = Univ.thy; |
|
1009 |
val rec_specs = |
|
1010 |
[("list", "univ(A)", |
|
1011 |
[(["Nil"], "i"), |
|
1012 |
(["Cons"], "[i,i]=>i")])]; |
|
1013 |
val rec_styp = "i=>i"; |
|
130 | 1014 |
val ext = None; |
103 | 1015 |
val sintrs = |
1016 |
["Nil : list(A)", |
|
1017 |
"[| a: A; l: list(A) |] ==> Cons(a,l) : list(A)"]; |
|
1018 |
val monos = []; |
|
130 | 1019 |
val type_intrs = datatype_intrs; |
103 | 1020 |
val type_elims = datatype_elims); |
1021 |
\end{ttbox} |
|
1022 |
\hrule |
|
1023 |
\caption{Defining the datatype of lists} \label{list-fig} |
|
1024 |
||
1025 |
\medskip |
|
1026 |
\begin{ttbox} |
|
130 | 1027 |
structure LList = CoDatatype_Fun |
103 | 1028 |
(val thy = QUniv.thy; |
1029 |
val rec_specs = |
|
1030 |
[("llist", "quniv(A)", |
|
1031 |
[(["LNil"], "i"), |
|
1032 |
(["LCons"], "[i,i]=>i")])]; |
|
1033 |
val rec_styp = "i=>i"; |
|
130 | 1034 |
val ext = None; |
103 | 1035 |
val sintrs = |
1036 |
["LNil : llist(A)", |
|
1037 |
"[| a: A; l: llist(A) |] ==> LCons(a,l) : llist(A)"]; |
|
1038 |
val monos = []; |
|
130 | 1039 |
val type_intrs = codatatype_intrs; |
1040 |
val type_elims = codatatype_elims); |
|
103 | 1041 |
\end{ttbox} |
1042 |
\hrule |
|
130 | 1043 |
\caption{Defining the codatatype of lazy lists} \label{llist-fig} |
103 | 1044 |
\end{figure} |
1045 |
||
1046 |
\subsection{Example: lists and lazy lists} |
|
1047 |
Figures \ref{list-fig} and~\ref{llist-fig} present the ML definitions of |
|
1048 |
lists and lazy lists, respectively. They highlight the (many) similarities |
|
130 | 1049 |
and (few) differences between datatype and codatatype definitions. |
103 | 1050 |
|
1051 |
Each form of list has two constructors, one for the empty list and one for |
|
1052 |
adding an element to a list. Each takes a parameter, defining the set of |
|
1053 |
lists over a given set~$A$. Each uses the appropriate domain from a |
|
1054 |
Isabelle/ZF theory: |
|
1055 |
\begin{itemize} |
|
1056 |
\item $\lst(A)$ specifies domain $\univ(A)$ and parent theory {\tt Univ.thy}. |
|
1057 |
||
1058 |
\item $\llist(A)$ specifies domain $\quniv(A)$ and parent theory {\tt |
|
1059 |
QUniv.thy}. |
|
1060 |
\end{itemize} |
|
1061 |
||
130 | 1062 |
Since $\lst(A)$ is a datatype, it enjoys a structural induction rule, {\tt |
1063 |
List.induct}: |
|
103 | 1064 |
\[ \infer{P(x)}{x\in\lst(A) & P(\Nil) |
1065 |
& \infer*{P(\Cons(a,l))}{[a\in A & l\in\lst(A) & P(l)]_{a,l}} } |
|
1066 |
\] |
|
1067 |
Induction and freeness yield the law $l\not=\Cons(a,l)$. To strengthen this, |
|
1068 |
Isabelle/ZF defines the rank of a set and proves that the standard pairs and |
|
1069 |
injections have greater rank than their components. An immediate consequence, |
|
1070 |
which justifies structural recursion on lists \cite[\S4.3]{paulson-set-II}, |
|
1071 |
is |
|
1072 |
\[ \rank(l) < \rank(\Cons(a,l)). \] |
|
1073 |
||
130 | 1074 |
Since $\llist(A)$ is a codatatype, it has no induction rule. Instead it has |
1075 |
the coinduction rule shown in \S\ref{coind-sec}. Since variant pairs and |
|
103 | 1076 |
injections are monotonic and need not have greater rank than their |
1077 |
components, fixedpoint operators can create cyclic constructions. For |
|
1078 |
example, the definition |
|
1079 |
\begin{eqnarray*} |
|
1080 |
\lconst(a) & \equiv & \lfp(\univ(a), \lambda l. \LCons(a,l)) |
|
1081 |
\end{eqnarray*} |
|
1082 |
yields $\lconst(a) = \LCons(a,\lconst(a))$. |
|
1083 |
||
1084 |
\medskip |
|
1085 |
It may be instructive to examine the definitions of the constructors and |
|
1086 |
case operator for $\lst(A)$. The definitions for $\llist(A)$ are similar. |
|
1087 |
The list constructors are defined as follows: |
|
1088 |
\begin{eqnarray*} |
|
1089 |
\Nil & = & \Inl(\emptyset) \\ |
|
1090 |
\Cons(a,l) & = & \Inr(\pair{a,l}) |
|
1091 |
\end{eqnarray*} |
|
1092 |
The operator $\lstcase$ performs case analysis on these two alternatives: |
|
1093 |
\begin{eqnarray*} |
|
1094 |
\lstcase(c,h) & \equiv & \case(\lambda u.c, \split(h)) |
|
1095 |
\end{eqnarray*} |
|
1096 |
Let us verify the two equations: |
|
1097 |
\begin{eqnarray*} |
|
1098 |
\lstcase(c, h, \Nil) & = & |
|
1099 |
\case(\lambda u.c, \split(h), \Inl(\emptyset)) \\ |
|
1100 |
& = & (\lambda u.c)(\emptyset) \\ |
|
130 | 1101 |
& = & c\\[1ex] |
103 | 1102 |
\lstcase(c, h, \Cons(x,y)) & = & |
1103 |
\case(\lambda u.c, \split(h), \Inr(\pair{x,y})) \\ |
|
1104 |
& = & \split(h, \pair{x,y}) \\ |
|
130 | 1105 |
& = & h(x,y) |
103 | 1106 |
\end{eqnarray*} |
1107 |
||
1108 |
\begin{figure} |
|
1109 |
\begin{small} |
|
1110 |
\begin{verbatim} |
|
1111 |
structure TF = Datatype_Fun |
|
1112 |
(val thy = Univ.thy; |
|
1113 |
val rec_specs = |
|
1114 |
[("tree", "univ(A)", |
|
1115 |
[(["Tcons"], "[i,i]=>i")]), |
|
1116 |
("forest", "univ(A)", |
|
1117 |
[(["Fnil"], "i"), |
|
1118 |
(["Fcons"], "[i,i]=>i")])]; |
|
1119 |
val rec_styp = "i=>i"; |
|
130 | 1120 |
val ext = None; |
103 | 1121 |
val sintrs = |
1122 |
["[| a:A; f: forest(A) |] ==> Tcons(a,f) : tree(A)", |
|
1123 |
"Fnil : forest(A)", |
|
1124 |
"[| t: tree(A); f: forest(A) |] ==> Fcons(t,f) : forest(A)"]; |
|
1125 |
val monos = []; |
|
130 | 1126 |
val type_intrs = datatype_intrs; |
103 | 1127 |
val type_elims = datatype_elims); |
1128 |
\end{verbatim} |
|
1129 |
\end{small} |
|
1130 |
\hrule |
|
1131 |
\caption{Defining the datatype of trees and forests} \label{tf-fig} |
|
1132 |
\end{figure} |
|
1133 |
||
1134 |
||
1135 |
\subsection{Example: mutual recursion} |
|
130 | 1136 |
In mutually recursive trees and forests~\cite[\S4.5]{paulson-set-II}, trees |
103 | 1137 |
have the one constructor $\Tcons$, while forests have the two constructors |
1138 |
$\Fnil$ and~$\Fcons$. Figure~\ref{tf-fig} presents the ML |
|
1139 |
definition. It has much in common with that of $\lst(A)$, including its |
|
1140 |
use of $\univ(A)$ for the domain and {\tt Univ.thy} for the parent theory. |
|
1141 |
The three introduction rules define the mutual recursion. The |
|
1142 |
distinguishing feature of this example is its two induction rules. |
|
1143 |
||
1144 |
The basic induction rule is called {\tt TF.induct}: |
|
1145 |
\[ \infer{P(x)}{x\in\TF(A) & |
|
1146 |
\infer*{P(\Tcons(a,f))} |
|
1147 |
{\left[\begin{array}{l} a\in A \\ |
|
1148 |
f\in\forest(A) \\ P(f) |
|
1149 |
\end{array} |
|
1150 |
\right]_{a,f}} |
|
1151 |
& P(\Fnil) |
|
130 | 1152 |
& \infer*{P(\Fcons(t,f))} |
103 | 1153 |
{\left[\begin{array}{l} t\in\tree(A) \\ P(t) \\ |
1154 |
f\in\forest(A) \\ P(f) |
|
1155 |
\end{array} |
|
1156 |
\right]_{t,f}} } |
|
1157 |
\] |
|
1158 |
This rule establishes a single predicate for $\TF(A)$, the union of the |
|
1159 |
recursive sets. |
|
1160 |
||
1161 |
Although such reasoning is sometimes useful |
|
1162 |
\cite[\S4.5]{paulson-set-II}, a proper mutual induction rule should establish |
|
1163 |
separate predicates for $\tree(A)$ and $\forest(A)$. The package calls this |
|
1164 |
rule {\tt TF.mutual\_induct}. Observe the usage of $P$ and $Q$ in the |
|
1165 |
induction hypotheses: |
|
1166 |
\[ \infer{(\forall z. z\in\tree(A)\imp P(z)) \conj |
|
1167 |
(\forall z. z\in\forest(A)\imp Q(z))} |
|
1168 |
{\infer*{P(\Tcons(a,f))} |
|
1169 |
{\left[\begin{array}{l} a\in A \\ |
|
1170 |
f\in\forest(A) \\ Q(f) |
|
1171 |
\end{array} |
|
1172 |
\right]_{a,f}} |
|
1173 |
& Q(\Fnil) |
|
130 | 1174 |
& \infer*{Q(\Fcons(t,f))} |
103 | 1175 |
{\left[\begin{array}{l} t\in\tree(A) \\ P(t) \\ |
1176 |
f\in\forest(A) \\ Q(f) |
|
1177 |
\end{array} |
|
1178 |
\right]_{t,f}} } |
|
1179 |
\] |
|
1180 |
As mentioned above, the package does not define a structural recursion |
|
1181 |
operator. I have described elsewhere how this is done |
|
1182 |
\cite[\S4.5]{paulson-set-II}. |
|
1183 |
||
1184 |
Both forest constructors have the form $\Inr(\cdots)$, |
|
1185 |
while the tree constructor has the form $\Inl(\cdots)$. This pattern would |
|
1186 |
hold regardless of how many tree or forest constructors there were. |
|
1187 |
\begin{eqnarray*} |
|
1188 |
\Tcons(a,l) & = & \Inl(\pair{a,l}) \\ |
|
1189 |
\Fnil & = & \Inr(\Inl(\emptyset)) \\ |
|
1190 |
\Fcons(a,l) & = & \Inr(\Inr(\pair{a,l})) |
|
1191 |
\end{eqnarray*} |
|
1192 |
There is only one case operator; it works on the union of the trees and |
|
1193 |
forests: |
|
1194 |
\begin{eqnarray*} |
|
1195 |
{\tt tree\_forest\_case}(f,c,g) & \equiv & |
|
1196 |
\case(\split(f),\, \case(\lambda u.c, \split(g))) |
|
1197 |
\end{eqnarray*} |
|
1198 |
||
1199 |
\begin{figure} |
|
1200 |
\begin{small} |
|
1201 |
\begin{verbatim} |
|
1202 |
structure Data = Datatype_Fun |
|
1203 |
(val thy = Univ.thy; |
|
1204 |
val rec_specs = |
|
1205 |
[("data", "univ(A Un B)", |
|
1206 |
[(["Con0"], "i"), |
|
1207 |
(["Con1"], "i=>i"), |
|
1208 |
(["Con2"], "[i,i]=>i"), |
|
1209 |
(["Con3"], "[i,i,i]=>i")])]; |
|
1210 |
val rec_styp = "[i,i]=>i"; |
|
130 | 1211 |
val ext = None; |
103 | 1212 |
val sintrs = |
1213 |
["Con0 : data(A,B)", |
|
1214 |
"[| a: A |] ==> Con1(a) : data(A,B)", |
|
1215 |
"[| a: A; b: B |] ==> Con2(a,b) : data(A,B)", |
|
1216 |
"[| a: A; b: B; d: data(A,B) |] ==> Con3(a,b,d) : data(A,B)"]; |
|
1217 |
val monos = []; |
|
130 | 1218 |
val type_intrs = datatype_intrs; |
103 | 1219 |
val type_elims = datatype_elims); |
1220 |
\end{verbatim} |
|
1221 |
\end{small} |
|
1222 |
\hrule |
|
1223 |
\caption{Defining the four-constructor sample datatype} \label{data-fig} |
|
1224 |
\end{figure} |
|
1225 |
||
1226 |
\subsection{A four-constructor datatype} |
|
1227 |
Finally let us consider a fairly general datatype. It has four |
|
130 | 1228 |
constructors $\Con_0$, \ldots, $\Con_3$, with the |
103 | 1229 |
corresponding arities. Figure~\ref{data-fig} presents the ML definition. |
1230 |
Because this datatype has two set parameters, $A$ and~$B$, it specifies |
|
1231 |
$\univ(A\un B)$ as its domain. The structural induction rule has four |
|
1232 |
minor premises, one per constructor: |
|
1233 |
\[ \infer{P(x)}{x\in\data(A,B) & |
|
1234 |
P(\Con_0) & |
|
1235 |
\infer*{P(\Con_1(a))}{[a\in A]_a} & |
|
1236 |
\infer*{P(\Con_2(a,b))} |
|
1237 |
{\left[\begin{array}{l} a\in A \\ b\in B \end{array} |
|
1238 |
\right]_{a,b}} & |
|
1239 |
\infer*{P(\Con_3(a,b,d))} |
|
1240 |
{\left[\begin{array}{l} a\in A \\ b\in B \\ |
|
1241 |
d\in\data(A,B) \\ P(d) |
|
1242 |
\end{array} |
|
1243 |
\right]_{a,b,d}} } |
|
1244 |
\] |
|
1245 |
||
1246 |
The constructor definitions are |
|
1247 |
\begin{eqnarray*} |
|
1248 |
\Con_0 & = & \Inl(\Inl(\emptyset)) \\ |
|
1249 |
\Con_1(a) & = & \Inl(\Inr(a)) \\ |
|
1250 |
\Con_2(a,b) & = & \Inr(\Inl(\pair{a,b})) \\ |
|
1251 |
\Con_3(a,b,c) & = & \Inr(\Inr(\pair{a,b,c})). |
|
1252 |
\end{eqnarray*} |
|
1253 |
The case operator is |
|
1254 |
\begin{eqnarray*} |
|
1255 |
{\tt data\_case}(f_0,f_1,f_2,f_3) & \equiv & |
|
1256 |
\case(\begin{array}[t]{@{}l} |
|
1257 |
\case(\lambda u.f_0,\; f_1),\, \\ |
|
1258 |
\case(\split(f_2),\; \split(\lambda v.\split(f_3(v)))) ) |
|
1259 |
\end{array} |
|
1260 |
\end{eqnarray*} |
|
1261 |
This may look cryptic, but the case equations are trivial to verify. |
|
1262 |
||
1263 |
In the constructor definitions, the injections are balanced. A more naive |
|
1264 |
approach is to define $\Con_3(a,b,c)$ as |
|
1265 |
$\Inr(\Inr(\Inr(\pair{a,b,c})))$; instead, each constructor has two |
|
1266 |
injections. The difference here is small. But the ZF examples include a |
|
1267 |
60-element enumeration type, where each constructor has 5 or~6 injections. |
|
1268 |
The naive approach would require 1 to~59 injections; the definitions would be |
|
1269 |
quadratic in size. It is like the difference between the binary and unary |
|
1270 |
numeral systems. |
|
1271 |
||
130 | 1272 |
The result structure contains the case operator and constructor definitions as |
1273 |
the theorem list \verb|con_defs|. It contains the case equations, such as |
|
103 | 1274 |
\begin{eqnarray*} |
1275 |
{\tt data\_case}(f_0,f_1,f_2,f_3,\Con_3(a,b,c)) & = & f_3(a,b,c), |
|
1276 |
\end{eqnarray*} |
|
1277 |
as the theorem list \verb|case_eqns|. There is one equation per constructor. |
|
1278 |
||
1279 |
\subsection{Proving freeness theorems} |
|
1280 |
There are two kinds of freeness theorems: |
|
1281 |
\begin{itemize} |
|
1282 |
\item {\bf injectiveness} theorems, such as |
|
1283 |
\[ \Con_2(a,b) = \Con_2(a',b') \bimp a=a' \conj b=b' \] |
|
1284 |
||
1285 |
\item {\bf distinctness} theorems, such as |
|
1286 |
\[ \Con_1(a) \not= \Con_2(a',b') \] |
|
1287 |
\end{itemize} |
|
1288 |
Since the number of such theorems is quadratic in the number of constructors, |
|
1289 |
the package does not attempt to prove them all. Instead it returns tools for |
|
1290 |
proving desired theorems --- either explicitly or `on the fly' during |
|
1291 |
simplification or classical reasoning. |
|
1292 |
||
1293 |
The theorem list \verb|free_iffs| enables the simplifier to perform freeness |
|
1294 |
reasoning. This works by incremental unfolding of constructors that appear in |
|
1295 |
equations. The theorem list contains logical equivalences such as |
|
1296 |
\begin{eqnarray*} |
|
1297 |
\Con_0=c & \bimp & c=\Inl(\Inl(\emptyset)) \\ |
|
1298 |
\Con_1(a)=c & \bimp & c=\Inl(\Inr(a)) \\ |
|
1299 |
& \vdots & \\ |
|
1300 |
\Inl(a)=\Inl(b) & \bimp & a=b \\ |
|
130 | 1301 |
\Inl(a)=\Inr(b) & \bimp & {\tt False} \\ |
103 | 1302 |
\pair{a,b} = \pair{a',b'} & \bimp & a=a' \conj b=b' |
1303 |
\end{eqnarray*} |
|
1304 |
For example, these rewrite $\Con_1(a)=\Con_1(b)$ to $a=b$ in four steps. |
|
1305 |
||
1306 |
The theorem list \verb|free_SEs| enables the classical |
|
1307 |
reasoner to perform similar replacements. It consists of elimination rules |
|
1308 |
to replace $\Con_0=c$ by $c=\Inl(\Inl(\emptyset))$, and so forth, in the |
|
1309 |
assumptions. |
|
1310 |
||
1311 |
Such incremental unfolding combines freeness reasoning with other proof |
|
1312 |
steps. It has the unfortunate side-effect of unfolding definitions of |
|
1313 |
constructors in contexts such as $\exists x.\Con_1(a)=x$, where they should |
|
1314 |
be left alone. Calling the Isabelle tactic {\tt fold\_tac con\_defs} |
|
1315 |
restores the defined constants. |
|
1316 |
\fi %CADE |
|
1317 |
||
1318 |
\section{Conclusions and future work} |
|
130 | 1319 |
The fixedpoint approach makes it easy to implement a powerful |
1320 |
package for inductive and coinductive definitions. It is efficient: it |
|
103 | 1321 |
processes most definitions in seconds and even a 60-constructor datatype |
1322 |
requires only two minutes. It is also simple: the package consists of |
|
1323 |
under 1100 lines (35K bytes) of Standard ML code. The first working |
|
1324 |
version took under a week to code. |
|
1325 |
||
1326 |
The approach is not restricted to set theory. It should be suitable for |
|
1327 |
any logic that has some notion of set and the Knaster-Tarski Theorem. |
|
1328 |
Indeed, Melham's inductive definition package for the HOL |
|
1329 |
system~\cite{camilleri92} implicitly proves this theorem. |
|
1330 |
||
130 | 1331 |
Datatype and codatatype definitions furthermore require a particular set |
103 | 1332 |
closed under a suitable notion of ordered pair. I intend to use the |
1333 |
Isabelle/ZF package as the basis for a higher-order logic one, using |
|
1334 |
Isabelle/HOL\@. The necessary theory is already |
|
130 | 1335 |
mechanized~\cite{paulson-coind}. HOL represents sets by unary predicates; |
103 | 1336 |
defining the corresponding types may cause complication. |
1337 |
||
1338 |
||
1339 |
\bibliographystyle{plain} |
|
1340 |
\bibliography{atp,theory,funprog,isabelle} |
|
1341 |
%%%%%\doendnotes |
|
1342 |
||
1343 |
\ifCADE\typeout{****Omitting appendices from CADE version!} |
|
1344 |
\else |
|
1345 |
\newpage |
|
1346 |
\appendix |
|
130 | 1347 |
\section{Inductive and coinductive definitions: users guide} |
1348 |
The ML functors \verb|Inductive_Fun| and \verb|CoInductive_Fun| build |
|
1349 |
inductive and coinductive definitions, respectively. This section describes |
|
103 | 1350 |
how to invoke them. |
1351 |
||
1352 |
\subsection{The result structure} |
|
1353 |
Many of the result structure's components have been discussed |
|
1354 |
in~\S\ref{basic-sec}; others are self-explanatory. |
|
1355 |
\begin{description} |
|
1356 |
\item[\tt thy] is the new theory containing the recursive sets. |
|
1357 |
||
1358 |
\item[\tt defs] is the list of definitions of the recursive sets. |
|
1359 |
||
1360 |
\item[\tt bnd\_mono] is a monotonicity theorem for the fixedpoint operator. |
|
1361 |
||
1362 |
\item[\tt unfold] is a fixedpoint equation for the recursive set (the union of |
|
1363 |
the recursive sets, in the case of mutual recursion). |
|
1364 |
||
1365 |
\item[\tt dom\_subset] is a theorem stating inclusion in the domain. |
|
1366 |
||
1367 |
\item[\tt intrs] is the list of introduction rules, now proved as theorems, for |
|
1368 |
the recursive sets. |
|
1369 |
||
1370 |
\item[\tt elim] is the elimination rule. |
|
1371 |
||
1372 |
\item[\tt mk\_cases] is a function to create simplified instances of {\tt |
|
1373 |
elim}, using freeness reasoning on some underlying datatype. |
|
1374 |
\end{description} |
|
1375 |
||
1376 |
For an inductive definition, the result structure contains two induction rules, |
|
130 | 1377 |
{\tt induct} and \verb|mutual_induct|. For a coinductive definition, it |
1378 |
contains the rule \verb|coinduct|. |
|
1379 |
||
1380 |
Figure~\ref{def-result-fig} summarizes the two result signatures, |
|
1381 |
specifying the types of all these components. |
|
103 | 1382 |
|
1383 |
\begin{figure} |
|
1384 |
\begin{ttbox} |
|
1385 |
sig |
|
1386 |
val thy : theory |
|
1387 |
val defs : thm list |
|
1388 |
val bnd_mono : thm |
|
1389 |
val unfold : thm |
|
1390 |
val dom_subset : thm |
|
1391 |
val intrs : thm list |
|
1392 |
val elim : thm |
|
1393 |
val mk_cases : thm list -> string -> thm |
|
1394 |
{\it(Inductive definitions only)} |
|
1395 |
val induct : thm |
|
1396 |
val mutual_induct: thm |
|
130 | 1397 |
{\it(Coinductive definitions only)} |
1398 |
val coinduct : thm |
|
103 | 1399 |
end |
1400 |
\end{ttbox} |
|
1401 |
\hrule |
|
130 | 1402 |
\caption{The result of a (co)inductive definition} \label{def-result-fig} |
103 | 1403 |
|
130 | 1404 |
\medskip |
103 | 1405 |
\begin{ttbox} |
1406 |
sig |
|
1407 |
val thy : theory |
|
1408 |
val rec_doms : (string*string) list |
|
1409 |
val sintrs : string list |
|
1410 |
val monos : thm list |
|
1411 |
val con_defs : thm list |
|
1412 |
val type_intrs : thm list |
|
1413 |
val type_elims : thm list |
|
1414 |
end |
|
1415 |
\end{ttbox} |
|
1416 |
\hrule |
|
130 | 1417 |
\caption{The argument of a (co)inductive definition} \label{def-arg-fig} |
103 | 1418 |
\end{figure} |
1419 |
||
1420 |
\subsection{The argument structure} |
|
130 | 1421 |
Both \verb|Inductive_Fun| and \verb|CoInductive_Fun| take the same argument |
103 | 1422 |
structure (Figure~\ref{def-arg-fig}). Its components are as follows: |
1423 |
\begin{description} |
|
1424 |
\item[\tt thy] is the definition's parent theory, which {\it must\/} |
|
1425 |
declare constants for the recursive sets. |
|
1426 |
||
1427 |
\item[\tt rec\_doms] is a list of pairs, associating the name of each recursive |
|
1428 |
set with its domain. |
|
1429 |
||
1430 |
\item[\tt sintrs] specifies the desired introduction rules as strings. |
|
1431 |
||
1432 |
\item[\tt monos] consists of monotonicity theorems for each operator applied |
|
1433 |
to a recursive set in the introduction rules. |
|
1434 |
||
1435 |
\item[\tt con\_defs] contains definitions of constants appearing in the |
|
130 | 1436 |
introduction rules. The (co)datatype package supplies the constructors' |
103 | 1437 |
definitions here. Most direct calls of \verb|Inductive_Fun| or |
130 | 1438 |
\verb|CoInductive_Fun| pass the empty list; one exception is the primitive |
103 | 1439 |
recursive functions example (\S\ref{primrec-sec}). |
1440 |
||
1441 |
\item[\tt type\_intrs] consists of introduction rules for type-checking the |
|
1442 |
definition, as discussed in~\S\ref{basic-sec}. They are applied using |
|
1443 |
depth-first search; you can trace the proof by setting |
|
1444 |
\verb|trace_DEPTH_FIRST := true|. |
|
1445 |
||
1446 |
\item[\tt type\_elims] consists of elimination rules for type-checking the |
|
1447 |
definition. They are presumed to be `safe' and are applied as much as |
|
1448 |
possible, prior to the {\tt type\_intrs} search. |
|
1449 |
\end{description} |
|
1450 |
The package has a few notable restrictions: |
|
1451 |
\begin{itemize} |
|
1452 |
\item The parent theory, {\tt thy}, must declare the recursive sets as |
|
1453 |
constants. You can extend a theory with new constants using {\tt |
|
1454 |
addconsts}, as illustrated in~\S\ref{ind-eg-sec}. If the inductive |
|
1455 |
definition also requires new concrete syntax, then it is simpler to |
|
1456 |
express the parent theory using a theory file. It is often convenient to |
|
1457 |
define an infix syntax for relations, say $a\prec b$ for $\pair{a,b}\in |
|
1458 |
R$. |
|
1459 |
||
1460 |
\item The names of the recursive sets must be identifiers, not infix |
|
1461 |
operators. |
|
1462 |
||
1463 |
\item Side-conditions must not be conjunctions. However, an introduction rule |
|
1464 |
may contain any number of side-conditions. |
|
1465 |
\end{itemize} |
|
1466 |
||
1467 |
||
130 | 1468 |
\section{Datatype and codatatype definitions: users guide} |
1469 |
The ML functors \verb|Datatype_Fun| and \verb|CoDatatype_Fun| define datatypes |
|
1470 |
and codatatypes, invoking \verb|Datatype_Fun| and |
|
1471 |
\verb|CoDatatype_Fun| to make the underlying (co)inductive definitions. |
|
103 | 1472 |
|
1473 |
||
1474 |
\subsection{The result structure} |
|
130 | 1475 |
The result structure extends that of (co)inductive definitions |
103 | 1476 |
(Figure~\ref{def-result-fig}) with several additional items: |
1477 |
\begin{ttbox} |
|
1478 |
val con_thy : theory |
|
1479 |
val con_defs : thm list |
|
1480 |
val case_eqns : thm list |
|
1481 |
val free_iffs : thm list |
|
1482 |
val free_SEs : thm list |
|
1483 |
val mk_free : string -> thm |
|
1484 |
\end{ttbox} |
|
1485 |
Most of these have been discussed in~\S\ref{data-sec}. Here is a summary: |
|
1486 |
\begin{description} |
|
1487 |
\item[\tt con\_thy] is a new theory containing definitions of the |
|
130 | 1488 |
(co)datatype's constructors and case operator. It also declares the |
103 | 1489 |
recursive sets as constants, so that it may serve as the parent |
130 | 1490 |
theory for the (co)inductive definition. |
103 | 1491 |
|
1492 |
\item[\tt con\_defs] is a list of definitions: the case operator followed by |
|
1493 |
the constructors. This theorem list can be supplied to \verb|mk_cases|, for |
|
1494 |
example. |
|
1495 |
||
1496 |
\item[\tt case\_eqns] is a list of equations, stating that the case operator |
|
1497 |
inverts each constructor. |
|
1498 |
||
1499 |
\item[\tt free\_iffs] is a list of logical equivalences to perform freeness |
|
1500 |
reasoning by rewriting. A typical application has the form |
|
1501 |
\begin{ttbox} |
|
1502 |
by (asm_simp_tac (ZF_ss addsimps free_iffs) 1); |
|
1503 |
\end{ttbox} |
|
1504 |
||
1505 |
\item[\tt free\_SEs] is a list of `safe' elimination rules to perform freeness |
|
1506 |
reasoning. It can be supplied to \verb|eresolve_tac| or to the classical |
|
1507 |
reasoner: |
|
1508 |
\begin{ttbox} |
|
1509 |
by (fast_tac (ZF_cs addSEs free_SEs) 1); |
|
1510 |
\end{ttbox} |
|
1511 |
||
1512 |
\item[\tt mk\_free] is a function to prove freeness properties, specified as |
|
1513 |
strings. The theorems can be expressed in various forms, such as logical |
|
1514 |
equivalences or elimination rules. |
|
1515 |
\end{description} |
|
1516 |
||
1517 |
The result structure also inherits everything from the underlying |
|
130 | 1518 |
(co)inductive definition, such as the introduction rules, elimination rule, |
179 | 1519 |
and (co)induction rule. |
103 | 1520 |
|
1521 |
||
1522 |
\begin{figure} |
|
1523 |
\begin{ttbox} |
|
1524 |
sig |
|
1525 |
val thy : theory |
|
1526 |
val rec_specs : (string * string * (string list*string)list) list |
|
1527 |
val rec_styp : string |
|
1528 |
val ext : Syntax.sext option |
|
1529 |
val sintrs : string list |
|
1530 |
val monos : thm list |
|
1531 |
val type_intrs: thm list |
|
1532 |
val type_elims: thm list |
|
1533 |
end |
|
1534 |
\end{ttbox} |
|
1535 |
\hrule |
|
130 | 1536 |
\caption{The argument of a (co)datatype definition} \label{data-arg-fig} |
103 | 1537 |
\end{figure} |
1538 |
||
1539 |
\subsection{The argument structure} |
|
130 | 1540 |
Both (co)datatype functors take the same argument structure |
1541 |
(Figure~\ref{data-arg-fig}). It does not extend that for (co)inductive |
|
103 | 1542 |
definitions, but shares several components and passes them uninterpreted to |
1543 |
\verb|Datatype_Fun| or |
|
130 | 1544 |
\verb|CoDatatype_Fun|. The new components are as follows: |
103 | 1545 |
\begin{description} |
130 | 1546 |
\item[\tt thy] is the (co)datatype's parent theory. It {\it must not\/} |
1547 |
declare constants for the recursive sets. Recall that (co)inductive |
|
103 | 1548 |
definitions have the opposite restriction. |
1549 |
||
1550 |
\item[\tt rec\_specs] is a list of triples of the form ({\it recursive set\/}, |
|
1551 |
{\it domain\/}, {\it constructors\/}) for each mutually recursive set. {\it |
|
1552 |
Constructors\/} is a list of the form (names, type). See the discussion and |
|
1553 |
examples in~\S\ref{data-sec}. |
|
1554 |
||
1555 |
\item[\tt rec\_styp] is the common meta-type of the mutually recursive sets, |
|
1556 |
specified as a string. They must all have the same type because all must |
|
1557 |
take the same parameters. |
|
1558 |
||
1559 |
\item[\tt ext] is an optional syntax extension, usually omitted by writing |
|
1560 |
{\tt None}. You can supply mixfix syntax for the constructors by supplying |
|
1561 |
\begin{ttbox} |
|
1562 |
Some (Syntax.simple_sext [{\it mixfix declarations\/}]) |
|
1563 |
\end{ttbox} |
|
1564 |
\end{description} |
|
1565 |
The choice of domain is usually simple. Isabelle/ZF defines the set |
|
1566 |
$\univ(A)$, which contains~$A$ and is closed under the standard Cartesian |
|
1567 |
products and disjoint sums \cite[\S4.2]{paulson-set-II}. In a typical |
|
1568 |
datatype definition with set parameters $A_1$, \ldots, $A_k$, a suitable |
|
1569 |
domain for all the recursive sets is $\univ(A_1\un\cdots\un A_k)$. For a |
|
130 | 1570 |
codatatype definition, the set |
103 | 1571 |
$\quniv(A)$ contains~$A$ and is closed under the variant Cartesian products |
130 | 1572 |
and disjoint sums; the appropriate domain is |
103 | 1573 |
$\quniv(A_1\un\cdots\un A_k)$. |
1574 |
||
1575 |
The {\tt sintrs} specify the introduction rules, which govern the recursive |
|
179 | 1576 |
structure of the datatype. Introduction rules may involve monotone |
1577 |
operators and side-conditions to express things that go beyond the usual |
|
1578 |
notion of datatype. The theorem lists {\tt monos}, {\tt type\_intrs} and |
|
1579 |
{\tt type\_elims} should contain precisely what is needed for the |
|
1580 |
underlying (co)inductive definition. Isabelle/ZF defines lists of |
|
1581 |
type-checking rules that can be supplied for the latter two components: |
|
103 | 1582 |
\begin{itemize} |
179 | 1583 |
\item {\tt datatype\_intrs} and {\tt datatype\_elims} are rules |
103 | 1584 |
for $\univ(A)$. |
179 | 1585 |
\item {\tt co\_datatype\_intrs} and {\tt co\_datatype\_elims} are |
103 | 1586 |
rules for $\quniv(A)$. |
1587 |
\end{itemize} |
|
1588 |
In typical definitions, these theorem lists need not be supplemented with |
|
1589 |
other theorems. |
|
1590 |
||
1591 |
The constructor definitions' right-hand sides can overlap. A |
|
1592 |
simple example is the datatype for the combinators, whose constructors are |
|
1593 |
\begin{eqnarray*} |
|
1594 |
{\tt K} & \equiv & \Inl(\emptyset) \\ |
|
1595 |
{\tt S} & \equiv & \Inr(\Inl(\emptyset)) \\ |
|
1596 |
p{\tt\#}q & \equiv & \Inr(\Inl(\pair{p,q})) |
|
1597 |
\end{eqnarray*} |
|
1598 |
Unlike in previous versions of Isabelle, \verb|fold_tac| now ensures that the |
|
1599 |
longest right-hand sides are folded first. |
|
1600 |
||
1601 |
\fi |
|
1602 |
\end{document} |