| author | blanchet |
| Tue, 19 Nov 2013 18:38:25 +0100 | |
| changeset 54504 | 096f7d452164 |
| parent 54494 | a220071f6708 |
| child 54537 | f37354a894a3 |
| permissions | -rw-r--r-- |
| 52792 | 1 |
(* Title: Doc/Datatypes/Datatypes.thy |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
| 53617 | 3 |
Author: Lorenz Panny, TU Muenchen |
4 |
Author: Andrei Popescu, TU Muenchen |
|
5 |
Author: Dmitriy Traytel, TU Muenchen |
|
| 52792 | 6 |
|
7 |
Tutorial for (co)datatype definitions with the new package. |
|
8 |
*) |
|
9 |
||
10 |
theory Datatypes |
|
| 54402 | 11 |
imports Setup "~~/src/HOL/Library/Simps_Case_Conv" |
| 52792 | 12 |
begin |
13 |
||
| 52828 | 14 |
section {* Introduction
|
15 |
\label{sec:introduction} *}
|
|
| 52792 | 16 |
|
17 |
text {*
|
|
| 53028 | 18 |
The 2013 edition of Isabelle introduced a new definitional package for freely |
19 |
generated datatypes and codatatypes. The datatype support is similar to that |
|
20 |
provided by the earlier package due to Berghofer and Wenzel |
|
21 |
\cite{Berghofer-Wenzel:1999:TPHOL}, documented in the Isar reference manual
|
|
| 53535 | 22 |
\cite{isabelle-isar-ref}; indeed, replacing the keyword \keyw{datatype} by
|
| 53028 | 23 |
@{command datatype_new} is usually all that is needed to port existing theories
|
24 |
to use the new package. |
|
| 52840 | 25 |
|
| 52841 | 26 |
Perhaps the main advantage of the new package is that it supports recursion |
| 53619 | 27 |
through a large class of non-datatypes, such as finite sets: |
| 52792 | 28 |
*} |
29 |
||
| 53644 | 30 |
datatype_new 'a tree\<^sub>f\<^sub>s = Node\<^sub>f\<^sub>s (lbl\<^sub>f\<^sub>s: 'a) (sub\<^sub>f\<^sub>s: "'a tree\<^sub>f\<^sub>s fset") |
| 52792 | 31 |
|
32 |
text {*
|
|
| 52794 | 33 |
\noindent |
| 53025 | 34 |
Another strong point is the support for local definitions: |
| 52792 | 35 |
*} |
36 |
||
| 52805 | 37 |
context linorder |
38 |
begin |
|
| 52841 | 39 |
datatype_new flag = Less | Eq | Greater |
| 52805 | 40 |
end |
| 52792 | 41 |
|
42 |
text {*
|
|
| 52794 | 43 |
\noindent |
| 54187 | 44 |
Furthermore, the package provides a lot of convenience, including automatically |
45 |
generated discriminators, selectors, and relators as well as a wealth of |
|
46 |
properties about them. |
|
47 |
||
48 |
In addition to inductive datatypes, the new package supports coinductive |
|
49 |
datatypes, or \emph{codatatypes}, which allow infinite values. For example, the
|
|
50 |
following command introduces the type of lazy lists, which comprises both finite |
|
51 |
and infinite values: |
|
| 52794 | 52 |
*} |
53 |
||
| 53623 | 54 |
(*<*) |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
55 |
locale early |
| 54072 | 56 |
locale late |
| 53623 | 57 |
(*>*) |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
58 |
codatatype (*<*)(in early) (*>*)'a llist = LNil | LCons 'a "'a llist" |
| 52794 | 59 |
|
60 |
text {*
|
|
61 |
\noindent |
|
| 52840 | 62 |
Mixed inductive--coinductive recursion is possible via nesting. Compare the |
| 53028 | 63 |
following four Rose tree examples: |
| 52794 | 64 |
*} |
65 |
||
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
66 |
datatype_new (*<*)(in early) (*>*)'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f 'a "'a tree\<^sub>f\<^sub>f list" |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
67 |
datatype_new (*<*)(in early) (*>*)'a tree\<^sub>f\<^sub>i = Node\<^sub>f\<^sub>i 'a "'a tree\<^sub>f\<^sub>i llist" |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
68 |
codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>f = Node\<^sub>i\<^sub>f 'a "'a tree\<^sub>i\<^sub>f list" |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
69 |
codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist" |
| 52792 | 70 |
|
| 52794 | 71 |
text {*
|
| 54187 | 72 |
The first two tree types allow only paths of finite length, whereas the last two |
73 |
allow infinite paths. Orthogonally, the nodes in the first and third types have |
|
74 |
finitely many direct subtrees, whereas those of the second and fourth may have |
|
75 |
infinite branching. |
|
| 52840 | 76 |
|
| 52794 | 77 |
To use the package, it is necessary to import the @{theory BNF} theory, which
|
| 53552 | 78 |
can be precompiled into the \texttt{HOL-BNF} image. The following commands show
|
| 52806 | 79 |
how to launch jEdit/PIDE with the image loaded and how to build the image |
80 |
without launching jEdit: |
|
| 52794 | 81 |
*} |
82 |
||
83 |
text {*
|
|
84 |
\noindent |
|
| 52806 | 85 |
\ \ \ \ \texttt{isabelle jedit -l HOL-BNF} \\
|
| 52827 | 86 |
\noindent |
| 52828 | 87 |
\hbox{}\ \ \ \ \texttt{isabelle build -b HOL-BNF}
|
| 52794 | 88 |
*} |
89 |
||
90 |
text {*
|
|
| 52805 | 91 |
The package, like its predecessor, fully adheres to the LCF philosophy |
92 |
\cite{mgordon79}: The characteristic theorems associated with the specified
|
|
93 |
(co)datatypes are derived rather than introduced axiomatically.% |
|
| 53543 | 94 |
\footnote{If the @{text quick_and_dirty} option is enabled, some of the
|
| 52841 | 95 |
internal constructions and most of the internal proof obligations are skipped.} |
| 52805 | 96 |
The package's metatheory is described in a pair of papers |
| 53552 | 97 |
\cite{traytel-et-al-2012,blanchette-et-al-wit}. The central notion is that of a
|
98 |
\emph{bounded natural functor} (BNF)---a well-behaved type constructor for which
|
|
99 |
nested (co)recursion is supported. |
|
| 52792 | 100 |
|
| 52794 | 101 |
This tutorial is organized as follows: |
102 |
||
| 52805 | 103 |
\begin{itemize}
|
| 52822 | 104 |
\setlength{\itemsep}{0pt}
|
105 |
||
| 52805 | 106 |
\item Section \ref{sec:defining-datatypes}, ``Defining Datatypes,''
|
| 52822 | 107 |
describes how to specify datatypes using the @{command datatype_new} command.
|
| 52805 | 108 |
|
| 53018 | 109 |
\item Section \ref{sec:defining-recursive-functions}, ``Defining Recursive
|
| 52805 | 110 |
Functions,'' describes how to specify recursive functions using |
| 53535 | 111 |
@{command primrec_new}, \keyw{fun}, and \keyw{function}.
|
| 52805 | 112 |
|
113 |
\item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,''
|
|
| 53829 | 114 |
describes how to specify codatatypes using the @{command codatatype} command.
|
| 52805 | 115 |
|
| 53018 | 116 |
\item Section \ref{sec:defining-corecursive-functions}, ``Defining Corecursive
|
| 52805 | 117 |
Functions,'' describes how to specify corecursive functions using the |
| 53826 | 118 |
@{command primcorec} and @{command primcorecursive} commands.
|
| 52794 | 119 |
|
| 53018 | 120 |
\item Section \ref{sec:registering-bounded-natural-functors}, ``Registering
|
| 53552 | 121 |
Bounded Natural Functors,'' explains how to use the @{command bnf} command
|
122 |
to register arbitrary type constructors as BNFs. |
|
| 52805 | 123 |
|
| 53552 | 124 |
\item Section |
| 53617 | 125 |
\ref{sec:deriving-destructors-and-theorems-for-free-constructors},
|
126 |
``Deriving Destructors and Theorems for Free Constructors,'' explains how to |
|
| 53552 | 127 |
use the command @{command wrap_free_constructors} to derive destructor constants
|
128 |
and theorems for freely generated types, as performed internally by @{command
|
|
| 53829 | 129 |
datatype_new} and @{command codatatype}.
|
| 52794 | 130 |
|
| 53617 | 131 |
%\item Section \ref{sec:standard-ml-interface}, ``Standard ML Interface,''
|
132 |
%describes the package's programmatic interface. |
|
| 52794 | 133 |
|
| 53617 | 134 |
%\item Section \ref{sec:interoperability}, ``Interoperability,''
|
135 |
%is concerned with the packages' interaction with other Isabelle packages and |
|
136 |
%tools, such as the code generator and the counterexample generators. |
|
| 52805 | 137 |
|
| 53617 | 138 |
%\item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and
|
139 |
%Limitations,'' concludes with known open issues at the time of writing. |
|
| 52805 | 140 |
\end{itemize}
|
| 52822 | 141 |
|
142 |
||
143 |
\newbox\boxA |
|
| 54185 | 144 |
\setbox\boxA=\hbox{\texttt{NOSPAM}}
|
145 |
||
146 |
\newcommand\authoremaili{\texttt{blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak
|
|
| 52822 | 147 |
in.\allowbreak tum.\allowbreak de}} |
| 54185 | 148 |
\newcommand\authoremailii{\texttt{lore{\color{white}NOSPAM}\kern-\wd\boxA{}nz.panny@\allowbreak
|
| 53028 | 149 |
\allowbreak tum.\allowbreak de}} |
| 54185 | 150 |
\newcommand\authoremailiii{\texttt{pope{\color{white}NOSPAM}\kern-\wd\boxA{}scua@\allowbreak
|
| 52822 | 151 |
in.\allowbreak tum.\allowbreak de}} |
| 54185 | 152 |
\newcommand\authoremailiv{\texttt{tray{\color{white}NOSPAM}\kern-\wd\boxA{}tel@\allowbreak
|
| 52822 | 153 |
in.\allowbreak tum.\allowbreak de}} |
154 |
||
| 53028 | 155 |
The commands @{command datatype_new} and @{command primrec_new} are expected to
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
156 |
replace \keyw{datatype} and \keyw{primrec} in a future release. Authors of new
|
| 53535 | 157 |
theories are encouraged to use the new commands, and maintainers of older |
| 53028 | 158 |
theories may want to consider upgrading. |
| 52841 | 159 |
|
160 |
Comments and bug reports concerning either the tool or this tutorial should be |
|
| 53028 | 161 |
directed to the authors at \authoremaili, \authoremailii, \authoremailiii, |
162 |
and \authoremailiv. |
|
| 52794 | 163 |
*} |
164 |
||
| 53491 | 165 |
|
| 52827 | 166 |
section {* Defining Datatypes
|
| 52805 | 167 |
\label{sec:defining-datatypes} *}
|
168 |
||
169 |
text {*
|
|
| 53646 | 170 |
Datatypes can be specified using the @{command datatype_new} command.
|
| 52805 | 171 |
*} |
| 52792 | 172 |
|
| 52824 | 173 |
|
| 53617 | 174 |
subsection {* Introductory Examples
|
175 |
\label{ssec:datatype-introductory-examples} *}
|
|
| 52794 | 176 |
|
| 53646 | 177 |
text {*
|
178 |
Datatypes are illustrated through concrete examples featuring different flavors |
|
179 |
of recursion. More examples can be found in the directory |
|
| 54185 | 180 |
\verb|~~/src/HOL/|\allowbreak\verb|BNF/Examples|. |
| 53646 | 181 |
*} |
| 53621 | 182 |
|
183 |
subsubsection {* Nonrecursive Types
|
|
184 |
\label{sssec:datatype-nonrecursive-types} *}
|
|
| 52794 | 185 |
|
| 52805 | 186 |
text {*
|
| 53028 | 187 |
Datatypes are introduced by specifying the desired names and argument types for |
| 53491 | 188 |
their constructors. \emph{Enumeration} types are the simplest form of datatype.
|
| 53028 | 189 |
All their constructors are nullary: |
| 52805 | 190 |
*} |
191 |
||
| 52828 | 192 |
datatype_new trool = Truue | Faalse | Perhaaps |
| 52805 | 193 |
|
194 |
text {*
|
|
| 53028 | 195 |
\noindent |
| 53491 | 196 |
Here, @{const Truue}, @{const Faalse}, and @{const Perhaaps} have the type @{typ trool}.
|
| 53028 | 197 |
|
198 |
Polymorphic types are possible, such as the following option type, modeled after |
|
199 |
its homologue from the @{theory Option} theory:
|
|
| 52805 | 200 |
*} |
201 |
||
| 53025 | 202 |
(*<*) |
203 |
hide_const None Some |
|
204 |
(*>*) |
|
205 |
datatype_new 'a option = None | Some 'a |
|
| 52805 | 206 |
|
207 |
text {*
|
|
| 53028 | 208 |
\noindent |
| 53491 | 209 |
The constructors are @{text "None :: 'a option"} and
|
210 |
@{text "Some :: 'a \<Rightarrow> 'a option"}.
|
|
| 53028 | 211 |
|
212 |
The next example has three type parameters: |
|
| 52805 | 213 |
*} |
214 |
||
215 |
datatype_new ('a, 'b, 'c) triple = Triple 'a 'b 'c
|
|
216 |
||
| 53028 | 217 |
text {*
|
218 |
\noindent |
|
219 |
The constructor is |
|
| 53491 | 220 |
@{text "Triple :: 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> ('a, 'b, 'c) triple"}.
|
| 53028 | 221 |
Unlike in Standard ML, curried constructors are supported. The uncurried variant |
222 |
is also possible: |
|
223 |
*} |
|
224 |
||
225 |
datatype_new ('a, 'b, 'c) triple\<^sub>u = Triple\<^sub>u "'a * 'b * 'c"
|
|
226 |
||
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
227 |
text {*
|
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
228 |
\noindent |
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
229 |
Occurrences of nonatomic types on the right-hand side of the equal sign must be |
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
230 |
enclosed in double quotes, as is customary in Isabelle. |
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
231 |
*} |
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
232 |
|
| 52824 | 233 |
|
| 53621 | 234 |
subsubsection {* Simple Recursion
|
235 |
\label{sssec:datatype-simple-recursion} *}
|
|
| 52794 | 236 |
|
| 52805 | 237 |
text {*
|
| 53491 | 238 |
Natural numbers are the simplest example of a recursive type: |
| 52805 | 239 |
*} |
240 |
||
| 53332 | 241 |
datatype_new nat = Zero | Suc nat |
| 52805 | 242 |
|
243 |
text {*
|
|
| 53491 | 244 |
\noindent |
| 54187 | 245 |
Lists were shown in the introduction. Terminated lists are a variant that |
246 |
stores a value of type @{typ 'b} at the very end:
|
|
| 52805 | 247 |
*} |
248 |
||
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
249 |
datatype_new (*<*)(in early) (*>*)('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist"
|
| 52805 | 250 |
|
| 52824 | 251 |
|
| 53621 | 252 |
subsubsection {* Mutual Recursion
|
253 |
\label{sssec:datatype-mutual-recursion} *}
|
|
| 52794 | 254 |
|
| 52805 | 255 |
text {*
|
| 53552 | 256 |
\emph{Mutually recursive} types are introduced simultaneously and may refer to
|
257 |
each other. The example below introduces a pair of types for even and odd |
|
258 |
natural numbers: |
|
| 52805 | 259 |
*} |
260 |
||
| 53623 | 261 |
datatype_new even_nat = Even_Zero | Even_Suc odd_nat |
262 |
and odd_nat = Odd_Suc even_nat |
|
| 52805 | 263 |
|
264 |
text {*
|
|
| 53491 | 265 |
\noindent |
266 |
Arithmetic expressions are defined via terms, terms via factors, and factors via |
|
267 |
expressions: |
|
| 52805 | 268 |
*} |
269 |
||
| 52841 | 270 |
datatype_new ('a, 'b) exp =
|
271 |
Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
|
|
| 52805 | 272 |
and ('a, 'b) trm =
|
| 52841 | 273 |
Factor "('a, 'b) fct" | Prod "('a, 'b) fct" "('a, 'b) trm"
|
274 |
and ('a, 'b) fct =
|
|
275 |
Const 'a | Var 'b | Expr "('a, 'b) exp"
|
|
| 52805 | 276 |
|
| 52824 | 277 |
|
| 53621 | 278 |
subsubsection {* Nested Recursion
|
279 |
\label{sssec:datatype-nested-recursion} *}
|
|
| 52794 | 280 |
|
| 52805 | 281 |
text {*
|
| 53491 | 282 |
\emph{Nested recursion} occurs when recursive occurrences of a type appear under
|
283 |
a type constructor. The introduction showed some examples of trees with nesting |
|
| 53675 | 284 |
through lists. A more complex example, that reuses our @{type option} type,
|
| 53491 | 285 |
follows: |
| 52805 | 286 |
*} |
287 |
||
| 52843 | 288 |
datatype_new 'a btree = |
| 53025 | 289 |
BNode 'a "'a btree option" "'a btree option" |
| 52805 | 290 |
|
291 |
text {*
|
|
| 53491 | 292 |
\noindent |
293 |
Not all nestings are admissible. For example, this command will fail: |
|
| 52805 | 294 |
*} |
295 |
||
| 54187 | 296 |
datatype_new 'a wrong = W1 | W2 (*<*)'a |
| 53542 | 297 |
typ (*>*)"'a wrong \<Rightarrow> 'a" |
| 52806 | 298 |
|
299 |
text {*
|
|
| 53491 | 300 |
\noindent |
301 |
The issue is that the function arrow @{text "\<Rightarrow>"} allows recursion
|
|
302 |
only through its right-hand side. This issue is inherited by polymorphic |
|
303 |
datatypes defined in terms of~@{text "\<Rightarrow>"}:
|
|
304 |
*} |
|
305 |
||
306 |
datatype_new ('a, 'b) fn = Fn "'a \<Rightarrow> 'b"
|
|
| 54187 | 307 |
datatype_new 'a also_wrong = W1 | W2 (*<*)'a |
| 53542 | 308 |
typ (*>*)"('a also_wrong, 'a) fn"
|
| 53491 | 309 |
|
310 |
text {*
|
|
311 |
\noindent |
|
| 53621 | 312 |
This is legal: |
313 |
*} |
|
314 |
||
315 |
datatype_new 'a ftree = FTLeaf 'a | FTNode "'a \<Rightarrow> 'a ftree" |
|
316 |
||
317 |
text {*
|
|
318 |
\noindent |
|
| 53491 | 319 |
In general, type constructors @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
|
320 |
allow recursion on a subset of their type arguments @{text 'a\<^sub>1}, \ldots,
|
|
321 |
@{text 'a\<^sub>m}. These type arguments are called \emph{live}; the remaining
|
|
322 |
type arguments are called \emph{dead}. In @{typ "'a \<Rightarrow> 'b"} and
|
|
323 |
@{typ "('a, 'b) fn"}, the type variable @{typ 'a} is dead and @{typ 'b} is live.
|
|
| 53552 | 324 |
|
| 53647 | 325 |
Type constructors must be registered as BNFs to have live arguments. This is |
326 |
done automatically for datatypes and codatatypes introduced by the @{command
|
|
| 53829 | 327 |
datatype_new} and @{command codatatype} commands.
|
| 53552 | 328 |
Section~\ref{sec:registering-bounded-natural-functors} explains how to register
|
329 |
arbitrary type constructors as BNFs. |
|
| 54187 | 330 |
|
331 |
Here is another example that fails: |
|
| 52806 | 332 |
*} |
333 |
||
| 54187 | 334 |
datatype_new 'a pow_list = PNil 'a (*<*)'a |
335 |
datatype_new 'a pow_list' = PNil' 'a (*>*)| PCons "('a * 'a) pow_list"
|
|
336 |
||
337 |
text {*
|
|
338 |
\noindent |
|
339 |
This one features a different flavor of nesting, where the recursive call in the |
|
340 |
type specification occurs around (rather than inside) another type constructor. |
|
341 |
*} |
|
342 |
||
343 |
subsubsection {* Auxiliary Constants and Properties
|
|
344 |
\label{sssec:datatype-auxiliary-constants-and-properties} *}
|
|
| 52805 | 345 |
|
346 |
text {*
|
|
| 53491 | 347 |
The @{command datatype_new} command introduces various constants in addition to
|
| 53617 | 348 |
the constructors. With each datatype are associated set functions, a map |
349 |
function, a relator, discriminators, and selectors, all of which can be given |
|
| 54187 | 350 |
custom names. In the example below, the familiar names @{text null}, @{text hd},
|
351 |
@{text tl}, @{text set}, @{text map}, and @{text list_all2}, override the
|
|
352 |
default names @{text is_Nil}, @{text un_Cons1}, @{text un_Cons2},
|
|
| 54491 | 353 |
@{text set_list}, @{text map_list}, and @{text rel_list}:
|
| 52822 | 354 |
*} |
355 |
||
| 52841 | 356 |
(*<*) |
357 |
no_translations |
|
358 |
"[x, xs]" == "x # [xs]" |
|
359 |
"[x]" == "x # []" |
|
360 |
||
361 |
no_notation |
|
362 |
Nil ("[]") and
|
|
363 |
Cons (infixr "#" 65) |
|
364 |
||
| 53543 | 365 |
hide_type list |
| 54494 | 366 |
hide_const Nil Cons hd tl set map list_all2 |
| 53025 | 367 |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
368 |
context early begin |
| 52841 | 369 |
(*>*) |
| 53025 | 370 |
datatype_new (set: 'a) list (map: map rel: list_all2) = |
| 52822 | 371 |
null: Nil (defaults tl: Nil) |
| 53025 | 372 |
| Cons (hd: 'a) (tl: "'a list") |
| 52822 | 373 |
|
374 |
text {*
|
|
375 |
\noindent |
|
| 54187 | 376 |
|
377 |
\begin{tabular}{@ {}ll@ {}}
|
|
378 |
Constructors: & |
|
379 |
@{text "Nil \<Colon> 'a list"} \\
|
|
380 |
& |
|
381 |
@{text "Cons \<Colon> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"} \\
|
|
382 |
Discriminator: & |
|
383 |
@{text "null \<Colon> 'a list \<Rightarrow> bool"} \\
|
|
384 |
Selectors: & |
|
385 |
@{text "hd \<Colon> 'a list \<Rightarrow> 'a"} \\
|
|
386 |
& |
|
387 |
@{text "tl \<Colon> 'a list \<Rightarrow> 'a list"} \\
|
|
388 |
Set function: & |
|
389 |
@{text "set \<Colon> 'a list \<Rightarrow> 'a set"} \\
|
|
390 |
Map function: & |
|
391 |
@{text "map \<Colon> ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \\
|
|
392 |
Relator: & |
|
393 |
@{text "list_all2 \<Colon> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool"}
|
|
394 |
\end{tabular}
|
|
395 |
||
396 |
The discriminator @{const null} and the selectors @{const hd} and @{const tl}
|
|
397 |
are characterized as follows: |
|
| 52822 | 398 |
% |
| 53025 | 399 |
\[@{thm list.collapse(1)[of xs, no_vars]}
|
400 |
\qquad @{thm list.collapse(2)[of xs, no_vars]}\]
|
|
| 52822 | 401 |
% |
| 54187 | 402 |
For two-constructor datatypes, a single discriminator constant is sufficient. |
403 |
The discriminator associated with @{const Cons} is simply
|
|
| 53491 | 404 |
@{term "\<lambda>xs. \<not> null xs"}.
|
| 52822 | 405 |
|
| 53553 | 406 |
The @{text defaults} clause following the @{const Nil} constructor specifies a
|
407 |
default value for selectors associated with other constructors. Here, it is used |
|
408 |
to ensure that the tail of the empty list is itself (instead of being left |
|
| 53535 | 409 |
unspecified). |
| 52822 | 410 |
|
| 53617 | 411 |
Because @{const Nil} is nullary, it is also possible to use
|
| 53491 | 412 |
@{term "\<lambda>xs. xs = Nil"} as a discriminator. This is specified by
|
| 53534 | 413 |
entering ``@{text "="}'' instead of the identifier @{const null}. Although this
|
| 53535 | 414 |
may look appealing, the mixture of constructors and selectors in the |
| 53534 | 415 |
characteristic theorems can lead Isabelle's automation to switch between the |
416 |
constructor and the destructor view in surprising ways. |
|
| 52822 | 417 |
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
418 |
The usual mixfix syntax annotations are available for both types and |
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
419 |
constructors. For example: |
| 52805 | 420 |
*} |
| 52794 | 421 |
|
| 53025 | 422 |
(*<*) |
423 |
end |
|
424 |
(*>*) |
|
| 53552 | 425 |
datatype_new ('a, 'b) prod (infixr "*" 20) = Pair 'a 'b
|
426 |
||
427 |
text {* \blankline *}
|
|
| 52822 | 428 |
|
| 52841 | 429 |
datatype_new (set: 'a) list (map: map rel: list_all2) = |
| 52822 | 430 |
null: Nil ("[]")
|
| 52841 | 431 |
| Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) |
432 |
||
433 |
text {*
|
|
| 53535 | 434 |
\noindent |
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
435 |
Incidentally, this is how the traditional syntax can be set up: |
| 52841 | 436 |
*} |
437 |
||
438 |
syntax "_list" :: "args \<Rightarrow> 'a list" ("[(_)]")
|
|
439 |
||
| 53552 | 440 |
text {* \blankline *}
|
441 |
||
| 52841 | 442 |
translations |
443 |
"[x, xs]" == "x # [xs]" |
|
444 |
"[x]" == "x # []" |
|
| 52822 | 445 |
|
| 52824 | 446 |
|
| 53617 | 447 |
subsection {* Command Syntax
|
448 |
\label{ssec:datatype-command-syntax} *}
|
|
449 |
||
450 |
||
| 53621 | 451 |
subsubsection {* \keyw{datatype\_new}
|
452 |
\label{sssec:datatype-new} *}
|
|
| 52794 | 453 |
|
| 52822 | 454 |
text {*
|
| 53829 | 455 |
\begin{matharray}{rcl}
|
456 |
@{command_def "datatype_new"} & : & @{text "local_theory \<rightarrow> local_theory"}
|
|
457 |
\end{matharray}
|
|
| 52822 | 458 |
|
| 52824 | 459 |
@{rail "
|
| 53829 | 460 |
@@{command datatype_new} target? @{syntax dt_options}? \\
|
| 52824 | 461 |
(@{syntax dt_name} '=' (@{syntax ctor} + '|') + @'and')
|
| 52828 | 462 |
; |
| 53623 | 463 |
@{syntax_def dt_options}: '(' (('no_discs_sels' | 'rep_compat') + ',') ')'
|
| 52824 | 464 |
"} |
465 |
||
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
466 |
The syntactic entity \synt{target} can be used to specify a local
|
| 53534 | 467 |
context---e.g., @{text "(in linorder)"}. It is documented in the Isar reference
|
468 |
manual \cite{isabelle-isar-ref}.
|
|
469 |
% |
|
470 |
The optional target is optionally followed by datatype-specific options: |
|
| 52822 | 471 |
|
| 52824 | 472 |
\begin{itemize}
|
473 |
\setlength{\itemsep}{0pt}
|
|
474 |
||
475 |
\item |
|
| 53623 | 476 |
The @{text "no_discs_sels"} option indicates that no discriminators or selectors
|
| 53543 | 477 |
should be generated. |
| 52822 | 478 |
|
| 52824 | 479 |
\item |
| 53644 | 480 |
The @{text "rep_compat"} option indicates that the generated names should
|
481 |
contain optional (and normally not displayed) ``@{text "new."}'' components to
|
|
482 |
prevent clashes with a later call to \keyw{rep\_datatype}. See
|
|
| 52824 | 483 |
Section~\ref{ssec:datatype-compatibility-issues} for details.
|
484 |
\end{itemize}
|
|
| 52822 | 485 |
|
| 52827 | 486 |
The left-hand sides of the datatype equations specify the name of the type to |
| 53534 | 487 |
define, its type parameters, and additional information: |
| 52822 | 488 |
|
| 52824 | 489 |
@{rail "
|
| 53534 | 490 |
@{syntax_def dt_name}: @{syntax tyargs}? name @{syntax map_rel}? mixfix?
|
| 52824 | 491 |
; |
| 53534 | 492 |
@{syntax_def tyargs}: typefree | '(' ((name ':')? typefree + ',') ')'
|
| 52824 | 493 |
; |
| 53534 | 494 |
@{syntax_def map_rel}: '(' ((('map' | 'rel') ':' name) +) ')'
|
| 52824 | 495 |
"} |
| 52822 | 496 |
|
| 52827 | 497 |
\noindent |
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
498 |
The syntactic entity \synt{name} denotes an identifier, \synt{typefree}
|
| 53534 | 499 |
denotes fixed type variable (@{typ 'a}, @{typ 'b}, \ldots), and \synt{mixfix}
|
500 |
denotes the usual parenthesized mixfix notation. They are documented in the Isar |
|
501 |
reference manual \cite{isabelle-isar-ref}.
|
|
| 52822 | 502 |
|
| 52827 | 503 |
The optional names preceding the type variables allow to override the default |
| 54491 | 504 |
names of the set functions (@{text set1_t}, \ldots, @{text setM_t}).
|
| 53647 | 505 |
Inside a mutually recursive specification, all defined datatypes must |
506 |
mention exactly the same type variables in the same order. |
|
| 52822 | 507 |
|
| 52824 | 508 |
@{rail "
|
| 53534 | 509 |
@{syntax_def ctor}: (name ':')? name (@{syntax ctor_arg} * ) \\
|
510 |
@{syntax dt_sel_defaults}? mixfix?
|
|
| 52824 | 511 |
"} |
512 |
||
| 53535 | 513 |
\medskip |
514 |
||
| 52827 | 515 |
\noindent |
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
516 |
The main constituents of a constructor specification are the name of the |
| 52827 | 517 |
constructor and the list of its argument types. An optional discriminator name |
| 53554 | 518 |
can be supplied at the front to override the default name |
519 |
(@{text t.is_C\<^sub>j}).
|
|
| 52822 | 520 |
|
| 52824 | 521 |
@{rail "
|
| 53534 | 522 |
@{syntax_def ctor_arg}: type | '(' name ':' type ')'
|
| 52827 | 523 |
"} |
524 |
||
| 53535 | 525 |
\medskip |
526 |
||
| 52827 | 527 |
\noindent |
528 |
In addition to the type of a constructor argument, it is possible to specify a |
|
529 |
name for the corresponding selector to override the default name |
|
| 53554 | 530 |
(@{text un_C\<^sub>ji}). The same selector names can be reused for several
|
531 |
constructors as long as they share the same type. |
|
| 52827 | 532 |
|
533 |
@{rail "
|
|
| 53621 | 534 |
@{syntax_def dt_sel_defaults}: '(' 'defaults' (name ':' term +) ')'
|
| 52824 | 535 |
"} |
| 52827 | 536 |
|
537 |
\noindent |
|
538 |
Given a constructor |
|
539 |
@{text "C \<Colon> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<sigma>\<^sub>p \<Rightarrow> \<sigma>"},
|
|
540 |
default values can be specified for any selector |
|
541 |
@{text "un_D \<Colon> \<sigma> \<Rightarrow> \<tau>"}
|
|
| 53535 | 542 |
associated with other constructors. The specified default value must be of type |
| 52828 | 543 |
@{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<sigma>\<^sub>p \<Rightarrow> \<tau>"}
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
544 |
(i.e., it may depend on @{text C}'s arguments).
|
| 52822 | 545 |
*} |
546 |
||
| 53617 | 547 |
|
| 53621 | 548 |
subsubsection {* \keyw{datatype\_new\_compat}
|
549 |
\label{sssec:datatype-new-compat} *}
|
|
| 53617 | 550 |
|
551 |
text {*
|
|
| 53829 | 552 |
\begin{matharray}{rcl}
|
553 |
@{command_def "datatype_new_compat"} & : & @{text "local_theory \<rightarrow> local_theory"}
|
|
554 |
\end{matharray}
|
|
555 |
||
556 |
@{rail "
|
|
557 |
@@{command datatype_new_compat} names
|
|
558 |
"} |
|
559 |
||
560 |
\noindent |
|
| 53621 | 561 |
The old datatype package provides some functionality that is not yet replicated |
562 |
in the new package: |
|
563 |
||
564 |
\begin{itemize}
|
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
565 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
566 |
|
| 53621 | 567 |
\item It is integrated with \keyw{fun} and \keyw{function}
|
568 |
\cite{isabelle-function}, Nitpick \cite{isabelle-nitpick}, Quickcheck,
|
|
569 |
and other packages. |
|
570 |
||
571 |
\item It is extended by various add-ons, notably to produce instances of the |
|
572 |
@{const size} function.
|
|
573 |
\end{itemize}
|
|
574 |
||
575 |
\noindent |
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
576 |
New-style datatypes can in most cases be registered as old-style datatypes using |
| 53829 | 577 |
@{command datatype_new_compat}. The \textit{names} argument is a space-separated
|
578 |
list of type names that are mutually recursive. For example: |
|
| 53621 | 579 |
*} |
580 |
||
| 53623 | 581 |
datatype_new_compat even_nat odd_nat |
| 53621 | 582 |
|
583 |
text {* \blankline *}
|
|
584 |
||
| 53623 | 585 |
thm even_nat_odd_nat.size |
| 53621 | 586 |
|
587 |
text {* \blankline *}
|
|
588 |
||
| 53623 | 589 |
ML {* Datatype_Data.get_info @{theory} @{type_name even_nat} *}
|
| 53621 | 590 |
|
591 |
text {*
|
|
| 53748 | 592 |
A few remarks concern nested recursive datatypes only: |
593 |
||
594 |
\begin{itemize}
|
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
595 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
596 |
|
| 53748 | 597 |
\item The old-style, nested-as-mutual induction rule, iterator theorems, and |
598 |
recursor theorems are generated under their usual names but with ``@{text
|
|
599 |
"compat_"}'' prefixed (e.g., @{text compat_tree.induct}).
|
|
600 |
||
601 |
\item All types through which recursion takes place must be new-style datatypes |
|
602 |
or the function type. In principle, it should be possible to support old-style |
|
603 |
datatypes as well, but the command does not support this yet (and there is |
|
604 |
currently no way to register old-style datatypes as new-style datatypes). |
|
| 54184 | 605 |
|
606 |
\item The recursor produced for types that recurse through functions has a |
|
607 |
different signature than with the old package. This makes it impossible to use |
|
608 |
the old \keyw{primrec} command.
|
|
| 53748 | 609 |
\end{itemize}
|
610 |
||
611 |
An alternative to @{command datatype_new_compat} is to use the old package's
|
|
612 |
\keyw{rep\_datatype} command. The associated proof obligations must then be
|
|
613 |
discharged manually. |
|
| 53617 | 614 |
*} |
615 |
||
616 |
||
617 |
subsection {* Generated Constants
|
|
618 |
\label{ssec:datatype-generated-constants} *}
|
|
619 |
||
620 |
text {*
|
|
| 53623 | 621 |
Given a datatype @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
|
| 53617 | 622 |
with $m > 0$ live type variables and $n$ constructors |
623 |
@{text "t.C\<^sub>1"}, \ldots, @{text "t.C\<^sub>n"}, the
|
|
624 |
following auxiliary constants are introduced: |
|
625 |
||
626 |
\begin{itemize}
|
|
627 |
\setlength{\itemsep}{0pt}
|
|
628 |
||
| 54494 | 629 |
\item \relax{Case combinator}: @{text t.case_t} (rendered using the familiar
|
| 53617 | 630 |
@{text case}--@{text of} syntax)
|
631 |
||
632 |
\item \relax{Discriminators}: @{text "t.is_C\<^sub>1"}, \ldots,
|
|
633 |
@{text "t.is_C\<^sub>n"}
|
|
634 |
||
635 |
\item \relax{Selectors}:
|
|
636 |
@{text t.un_C\<^sub>11}$, \ldots, @{text t.un_C\<^sub>1k\<^sub>1}, \\
|
|
637 |
\phantom{\relax{Selectors:}} \quad\vdots \\
|
|
638 |
\phantom{\relax{Selectors:}} @{text t.un_C\<^sub>n1}$, \ldots, @{text t.un_C\<^sub>nk\<^sub>n}.
|
|
639 |
||
640 |
\item \relax{Set functions} (or \relax{natural transformations}):
|
|
| 54491 | 641 |
@{text set1_t}, \ldots, @{text t.setm_t}
|
642 |
||
643 |
\item \relax{Map function} (or \relax{functorial action}): @{text t.map_t}
|
|
644 |
||
645 |
\item \relax{Relator}: @{text t.rel_t}
|
|
646 |
||
| 54494 | 647 |
\item \relax{Iterator}: @{text t.fold_t}
|
648 |
||
649 |
\item \relax{Recursor}: @{text t.rec_t}
|
|
| 53617 | 650 |
|
651 |
\end{itemize}
|
|
652 |
||
653 |
\noindent |
|
654 |
The case combinator, discriminators, and selectors are collectively called |
|
655 |
\emph{destructors}. The prefix ``@{text "t."}'' is an optional component of the
|
|
| 54491 | 656 |
names and is normally hidden. |
| 53617 | 657 |
*} |
658 |
||
659 |
||
| 52840 | 660 |
subsection {* Generated Theorems
|
661 |
\label{ssec:datatype-generated-theorems} *}
|
|
| 52828 | 662 |
|
663 |
text {*
|
|
| 53544 | 664 |
The characteristic theorems generated by @{command datatype_new} are grouped in
|
| 53623 | 665 |
three broad categories: |
| 53535 | 666 |
|
| 53543 | 667 |
\begin{itemize}
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
668 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
669 |
|
| 53543 | 670 |
\item The \emph{free constructor theorems} are properties about the constructors
|
671 |
and destructors that can be derived for any freely generated type. Internally, |
|
| 53542 | 672 |
the derivation is performed by @{command wrap_free_constructors}.
|
| 53535 | 673 |
|
| 53552 | 674 |
\item The \emph{functorial theorems} are properties of datatypes related to
|
675 |
their BNF nature. |
|
676 |
||
677 |
\item The \emph{inductive theorems} are properties of datatypes related to
|
|
| 53544 | 678 |
their inductive nature. |
| 53552 | 679 |
|
| 53543 | 680 |
\end{itemize}
|
| 53535 | 681 |
|
682 |
\noindent |
|
| 53542 | 683 |
The full list of named theorems can be obtained as usual by entering the |
| 53543 | 684 |
command \keyw{print\_theorems} immediately after the datatype definition.
|
| 53542 | 685 |
This list normally excludes low-level theorems that reveal internal |
| 53552 | 686 |
constructions. To make these accessible, add the line |
| 53542 | 687 |
*} |
| 53535 | 688 |
|
| 53542 | 689 |
declare [[bnf_note_all]] |
690 |
(*<*) |
|
691 |
declare [[bnf_note_all = false]] |
|
692 |
(*>*) |
|
| 53535 | 693 |
|
| 53552 | 694 |
text {*
|
695 |
\noindent |
|
696 |
to the top of the theory file. |
|
697 |
*} |
|
| 53535 | 698 |
|
| 53621 | 699 |
subsubsection {* Free Constructor Theorems
|
700 |
\label{sssec:free-constructor-theorems} *}
|
|
| 53535 | 701 |
|
| 53543 | 702 |
(*<*) |
| 53837 | 703 |
consts nonnull :: 'a |
| 53543 | 704 |
(*>*) |
705 |
||
| 53535 | 706 |
text {*
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
707 |
The first subgroup of properties is concerned with the constructors. |
| 53543 | 708 |
They are listed below for @{typ "'a list"}:
|
709 |
||
| 53552 | 710 |
\begin{indentblock}
|
| 53543 | 711 |
\begin{description}
|
| 53544 | 712 |
|
| 53642 | 713 |
\item[@{text "t."}\hthm{inject} @{text "[iff, induct_simp]"}\rm:] ~ \\
|
| 53544 | 714 |
@{thm list.inject[no_vars]}
|
715 |
||
| 53642 | 716 |
\item[@{text "t."}\hthm{distinct} @{text "[simp, induct_simp]"}\rm:] ~ \\
|
| 53543 | 717 |
@{thm list.distinct(1)[no_vars]} \\
|
718 |
@{thm list.distinct(2)[no_vars]}
|
|
719 |
||
| 53642 | 720 |
\item[@{text "t."}\hthm{exhaust} @{text "[cases t, case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
|
| 53543 | 721 |
@{thm list.exhaust[no_vars]}
|
722 |
||
| 53642 | 723 |
\item[@{text "t."}\hthm{nchotomy}\rm:] ~ \\
|
| 53543 | 724 |
@{thm list.nchotomy[no_vars]}
|
725 |
||
726 |
\end{description}
|
|
| 53552 | 727 |
\end{indentblock}
|
| 53543 | 728 |
|
729 |
\noindent |
|
| 53621 | 730 |
In addition, these nameless theorems are registered as safe elimination rules: |
731 |
||
732 |
\begin{indentblock}
|
|
733 |
\begin{description}
|
|
734 |
||
| 54386 | 735 |
\item[@{text "t."}\hthm{distinct {\upshape[}THEN notE}@{text ", elim!"}\hthm{\upshape]}\rm:] ~ \\
|
| 53621 | 736 |
@{thm list.distinct(1)[THEN notE, elim!, no_vars]} \\
|
737 |
@{thm list.distinct(2)[THEN notE, elim!, no_vars]}
|
|
738 |
||
739 |
\end{description}
|
|
740 |
\end{indentblock}
|
|
741 |
||
742 |
\noindent |
|
| 53543 | 743 |
The next subgroup is concerned with the case combinator: |
744 |
||
| 53552 | 745 |
\begin{indentblock}
|
| 53543 | 746 |
\begin{description}
|
| 53544 | 747 |
|
| 53798 | 748 |
\item[@{text "t."}\hthm{case} @{text "[simp, code]"}\rm:] ~ \\
|
| 53543 | 749 |
@{thm list.case(1)[no_vars]} \\
|
750 |
@{thm list.case(2)[no_vars]}
|
|
751 |
||
| 53642 | 752 |
\item[@{text "t."}\hthm{case\_cong}\rm:] ~ \\
|
| 53543 | 753 |
@{thm list.case_cong[no_vars]}
|
754 |
||
| 53642 | 755 |
\item[@{text "t."}\hthm{weak\_case\_cong} @{text "[cong]"}\rm:] ~ \\
|
| 53543 | 756 |
@{thm list.weak_case_cong[no_vars]}
|
757 |
||
| 53642 | 758 |
\item[@{text "t."}\hthm{split}\rm:] ~ \\
|
| 53543 | 759 |
@{thm list.split[no_vars]}
|
760 |
||
| 53642 | 761 |
\item[@{text "t."}\hthm{split\_asm}\rm:] ~ \\
|
| 53543 | 762 |
@{thm list.split_asm[no_vars]}
|
763 |
||
| 53544 | 764 |
\item[@{text "t."}\hthm{splits} = @{text "split split_asm"}]
|
| 53543 | 765 |
|
766 |
\end{description}
|
|
| 53552 | 767 |
\end{indentblock}
|
| 53543 | 768 |
|
769 |
\noindent |
|
770 |
The third and last subgroup revolves around discriminators and selectors: |
|
771 |
||
| 53552 | 772 |
\begin{indentblock}
|
| 53543 | 773 |
\begin{description}
|
| 53544 | 774 |
|
| 53694 | 775 |
\item[@{text "t."}\hthm{disc} @{text "[simp]"}\rm:] ~ \\
|
776 |
@{thm list.disc(1)[no_vars]} \\
|
|
777 |
@{thm list.disc(2)[no_vars]}
|
|
778 |
||
| 53703 | 779 |
\item[@{text "t."}\hthm{discI}\rm:] ~ \\
|
780 |
@{thm list.discI(1)[no_vars]} \\
|
|
781 |
@{thm list.discI(2)[no_vars]}
|
|
782 |
||
| 53805 | 783 |
\item[@{text "t."}\hthm{sel} @{text "[simp, code]"}\rm:] ~ \\
|
| 53694 | 784 |
@{thm list.sel(1)[no_vars]} \\
|
785 |
@{thm list.sel(2)[no_vars]}
|
|
| 53543 | 786 |
|
| 53642 | 787 |
\item[@{text "t."}\hthm{collapse} @{text "[simp]"}\rm:] ~ \\
|
| 53543 | 788 |
@{thm list.collapse(1)[no_vars]} \\
|
789 |
@{thm list.collapse(2)[no_vars]}
|
|
790 |
||
| 53837 | 791 |
\item[@{text "t."}\hthm{disc\_exclude} @{text "[dest]"}\rm:] ~ \\
|
| 53543 | 792 |
These properties are missing for @{typ "'a list"} because there is only one
|
793 |
proper discriminator. Had the datatype been introduced with a second |
|
| 53837 | 794 |
discriminator called @{const nonnull}, they would have read thusly: \\[\jot]
|
795 |
@{prop "null list \<Longrightarrow> \<not> nonnull list"} \\
|
|
796 |
@{prop "nonnull list \<Longrightarrow> \<not> null list"}
|
|
| 53543 | 797 |
|
| 53642 | 798 |
\item[@{text "t."}\hthm{disc\_exhaust} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
|
| 53543 | 799 |
@{thm list.disc_exhaust[no_vars]}
|
800 |
||
| 53916 | 801 |
\item[@{text "t."}\hthm{sel\_exhaust} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
|
802 |
@{thm list.sel_exhaust[no_vars]}
|
|
803 |
||
| 53642 | 804 |
\item[@{text "t."}\hthm{expand}\rm:] ~ \\
|
| 53543 | 805 |
@{thm list.expand[no_vars]}
|
806 |
||
| 53917 | 807 |
\item[@{text "t."}\hthm{sel\_split}\rm:] ~ \\
|
808 |
@{thm list.sel_split[no_vars]}
|
|
809 |
||
810 |
\item[@{text "t."}\hthm{sel\_split\_asm}\rm:] ~ \\
|
|
811 |
@{thm list.sel_split_asm[no_vars]}
|
|
812 |
||
| 54491 | 813 |
\item[@{text "t."}\hthm{case\_eq\_if}\rm:] ~ \\
|
814 |
@{thm list.case_eq_if[no_vars]}
|
|
| 53543 | 815 |
|
816 |
\end{description}
|
|
| 53552 | 817 |
\end{indentblock}
|
| 54152 | 818 |
|
819 |
\noindent |
|
820 |
In addition, equational versions of @{text t.disc} are registered with the @{text "[code]"}
|
|
821 |
attribute. |
|
| 53552 | 822 |
*} |
823 |
||
824 |
||
| 53621 | 825 |
subsubsection {* Functorial Theorems
|
826 |
\label{sssec:functorial-theorems} *}
|
|
| 53552 | 827 |
|
828 |
text {*
|
|
| 53623 | 829 |
The BNF-related theorem are as follows: |
| 53552 | 830 |
|
831 |
\begin{indentblock}
|
|
832 |
\begin{description}
|
|
833 |
||
| 53798 | 834 |
\item[@{text "t."}\hthm{set} @{text "[simp, code]"}\rm:] ~ \\
|
| 53694 | 835 |
@{thm list.set(1)[no_vars]} \\
|
836 |
@{thm list.set(2)[no_vars]}
|
|
| 53552 | 837 |
|
| 53798 | 838 |
\item[@{text "t."}\hthm{map} @{text "[simp, code]"}\rm:] ~ \\
|
| 53552 | 839 |
@{thm list.map(1)[no_vars]} \\
|
840 |
@{thm list.map(2)[no_vars]}
|
|
841 |
||
| 54146 | 842 |
\item[@{text "t."}\hthm{rel\_inject} @{text "[simp]"}\rm:] ~ \\
|
| 53552 | 843 |
@{thm list.rel_inject(1)[no_vars]} \\
|
844 |
@{thm list.rel_inject(2)[no_vars]}
|
|
845 |
||
| 54146 | 846 |
\item[@{text "t."}\hthm{rel\_distinct} @{text "[simp]"}\rm:] ~ \\
|
| 53552 | 847 |
@{thm list.rel_distinct(1)[no_vars]} \\
|
848 |
@{thm list.rel_distinct(2)[no_vars]}
|
|
849 |
||
850 |
\end{description}
|
|
851 |
\end{indentblock}
|
|
| 54146 | 852 |
|
853 |
\noindent |
|
854 |
In addition, equational versions of @{text t.rel_inject} and @{text
|
|
855 |
rel_distinct} are registered with the @{text "[code]"} attribute.
|
|
| 53535 | 856 |
*} |
857 |
||
858 |
||
| 53621 | 859 |
subsubsection {* Inductive Theorems
|
860 |
\label{sssec:inductive-theorems} *}
|
|
| 53535 | 861 |
|
862 |
text {*
|
|
| 53623 | 863 |
The inductive theorems are as follows: |
| 53544 | 864 |
|
| 53552 | 865 |
\begin{indentblock}
|
| 53544 | 866 |
\begin{description}
|
867 |
||
| 53642 | 868 |
\item[@{text "t."}\hthm{induct} @{text "[induct t, case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
|
| 53544 | 869 |
@{thm list.induct[no_vars]}
|
870 |
||
| 53642 | 871 |
\item[@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
|
| 53544 | 872 |
Given $m > 1$ mutually recursive datatypes, this induction rule can be used to |
873 |
prove $m$ properties simultaneously. |
|
| 52828 | 874 |
|
| 53798 | 875 |
\item[@{text "t."}\hthm{fold} @{text "[simp, code]"}\rm:] ~ \\
|
| 53544 | 876 |
@{thm list.fold(1)[no_vars]} \\
|
877 |
@{thm list.fold(2)[no_vars]}
|
|
878 |
||
| 53798 | 879 |
\item[@{text "t."}\hthm{rec} @{text "[simp, code]"}\rm:] ~ \\
|
| 53544 | 880 |
@{thm list.rec(1)[no_vars]} \\
|
881 |
@{thm list.rec(2)[no_vars]}
|
|
882 |
||
883 |
\end{description}
|
|
| 53552 | 884 |
\end{indentblock}
|
| 53544 | 885 |
|
886 |
\noindent |
|
887 |
For convenience, @{command datatype_new} also provides the following collection:
|
|
888 |
||
| 53552 | 889 |
\begin{indentblock}
|
| 53544 | 890 |
\begin{description}
|
891 |
||
892 |
\item[@{text "t."}\hthm{simps} = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.rec} @{text t.fold} @{text t.map} @{text t.rel_inject}] ~ \\
|
|
| 53694 | 893 |
@{text t.rel_distinct} @{text t.set}
|
| 53544 | 894 |
|
895 |
\end{description}
|
|
| 53552 | 896 |
\end{indentblock}
|
| 52828 | 897 |
*} |
898 |
||
| 52794 | 899 |
|
| 52827 | 900 |
subsection {* Compatibility Issues
|
| 52824 | 901 |
\label{ssec:datatype-compatibility-issues} *}
|
| 52794 | 902 |
|
| 52828 | 903 |
text {*
|
| 53997 | 904 |
The command @{command datatype_new} has been designed to be highly compatible
|
905 |
with the old \keyw{datatype}, to ease migration. There are nonetheless a few
|
|
| 53647 | 906 |
incompatibilities that may arise when porting to the new package: |
907 |
||
908 |
\begin{itemize}
|
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
909 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
910 |
|
| 53647 | 911 |
\item \emph{The Standard ML interfaces are different.} Tools and extensions
|
912 |
written to call the old ML interfaces will need to be adapted to the new |
|
913 |
interfaces. Little has been done so far in this direction. Whenever possible, it |
|
914 |
is recommended to use @{command datatype_new_compat} or \keyw{rep\_datatype}
|
|
915 |
to register new-style datatypes as old-style datatypes. |
|
916 |
||
| 54494 | 917 |
\item \emph{The constants @{text "t_case"} and @{text "t_rec"} are now called
|
918 |
@{text "case_t"} and @{text "rec_t"}.
|
|
919 |
||
920 |
\item \emph{The recursor @{text "rec_t"} has a different signature for nested
|
|
| 54185 | 921 |
recursive datatypes.} In the old package, nested recursion through non-functions |
922 |
was internally reduced to mutual recursion. This reduction was visible in the |
|
923 |
type of the recursor, used by \keyw{primrec}. Recursion through functions was
|
|
924 |
handled specially. In the new package, nested recursion (for functions and |
|
925 |
non-functions) is handled in a more modular fashion. The old-style recursor can |
|
926 |
be generated on demand using @{command primrec_new}, as explained in
|
|
| 53647 | 927 |
Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
|
928 |
new-style datatypes. |
|
929 |
||
| 54287 | 930 |
\item \emph{Accordingly, the induction rule is different for nested recursive
|
931 |
datatypes.} Again, the old-style induction rule can be generated on demand using |
|
932 |
@{command primrec_new}, as explained in
|
|
| 53647 | 933 |
Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
|
934 |
new-style datatypes. |
|
| 52828 | 935 |
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
936 |
\item \emph{The internal constructions are completely different.} Proof texts
|
| 53647 | 937 |
that unfold the definition of constants introduced by \keyw{datatype} will be
|
938 |
difficult to port. |
|
939 |
||
940 |
\item \emph{A few theorems have different names.}
|
|
| 53997 | 941 |
The properties @{text t.cases} and @{text t.recs} have been renamed
|
| 53647 | 942 |
@{text t.case} and @{text t.rec}. For non-mutually recursive datatypes,
|
943 |
@{text t.inducts} is available as @{text t.induct}.
|
|
944 |
For $m > 1$ mutually recursive datatypes, |
|
| 53997 | 945 |
@{text "t\<^sub>1_\<dots>_t\<^sub>m.inducts(i)"} has been renamed
|
| 53647 | 946 |
@{text "t\<^sub>i.induct"}.
|
947 |
||
948 |
\item \emph{The @{text t.simps} collection has been extended.}
|
|
949 |
Previously available theorems are available at the same index. |
|
950 |
||
951 |
\item \emph{Variables in generated properties have different names.} This is
|
|
952 |
rarely an issue, except in proof texts that refer to variable names in the |
|
953 |
@{text "[where \<dots>]"} attribute. The solution is to use the more robust
|
|
954 |
@{text "[of \<dots>]"} syntax.
|
|
955 |
\end{itemize}
|
|
956 |
||
957 |
In the other direction, there is currently no way to register old-style |
|
958 |
datatypes as new-style datatypes. If the goal is to define new-style datatypes |
|
959 |
with nested recursion through old-style datatypes, the old-style |
|
960 |
datatypes can be registered as a BNF |
|
961 |
(Section~\ref{sec:registering-bounded-natural-functors}). If the goal is
|
|
962 |
to derive discriminators and selectors, this can be achieved using @{command
|
|
963 |
wrap_free_constructors} |
|
964 |
(Section~\ref{sec:deriving-destructors-and-theorems-for-free-constructors}).
|
|
| 52828 | 965 |
*} |
966 |
||
| 52792 | 967 |
|
| 52827 | 968 |
section {* Defining Recursive Functions
|
| 52805 | 969 |
\label{sec:defining-recursive-functions} *}
|
970 |
||
971 |
text {*
|
|
| 54183 | 972 |
Recursive functions over datatypes can be specified using the @{command
|
973 |
primrec_new} command, which supports primitive recursion, or using the more |
|
974 |
general \keyw{fun} and \keyw{function} commands. Here, the focus is on @{command
|
|
| 53644 | 975 |
primrec_new}; the other two commands are described in a separate tutorial |
| 53646 | 976 |
\cite{isabelle-function}.
|
| 52828 | 977 |
|
| 53621 | 978 |
%%% TODO: partial_function |
| 52805 | 979 |
*} |
| 52792 | 980 |
|
| 52805 | 981 |
|
| 53617 | 982 |
subsection {* Introductory Examples
|
983 |
\label{ssec:primrec-introductory-examples} *}
|
|
| 52828 | 984 |
|
| 53646 | 985 |
text {*
|
986 |
Primitive recursion is illustrated through concrete examples based on the |
|
987 |
datatypes defined in Section~\ref{ssec:datatype-introductory-examples}. More
|
|
988 |
examples can be found in the directory \verb|~~/src/HOL/BNF/Examples|. |
|
989 |
*} |
|
990 |
||
| 53621 | 991 |
|
992 |
subsubsection {* Nonrecursive Types
|
|
993 |
\label{sssec:primrec-nonrecursive-types} *}
|
|
| 52828 | 994 |
|
| 52841 | 995 |
text {*
|
| 53621 | 996 |
Primitive recursion removes one layer of constructors on the left-hand side in |
997 |
each equation. For example: |
|
| 52841 | 998 |
*} |
999 |
||
1000 |
primrec_new bool_of_trool :: "trool \<Rightarrow> bool" where |
|
| 53621 | 1001 |
"bool_of_trool Faalse \<longleftrightarrow> False" | |
1002 |
"bool_of_trool Truue \<longleftrightarrow> True" |
|
| 52841 | 1003 |
|
| 53621 | 1004 |
text {* \blankline *}
|
| 52841 | 1005 |
|
| 53025 | 1006 |
primrec_new the_list :: "'a option \<Rightarrow> 'a list" where |
1007 |
"the_list None = []" | |
|
1008 |
"the_list (Some a) = [a]" |
|
| 52841 | 1009 |
|
| 53621 | 1010 |
text {* \blankline *}
|
1011 |
||
| 53025 | 1012 |
primrec_new the_default :: "'a \<Rightarrow> 'a option \<Rightarrow> 'a" where |
1013 |
"the_default d None = d" | |
|
1014 |
"the_default _ (Some a) = a" |
|
| 52843 | 1015 |
|
| 53621 | 1016 |
text {* \blankline *}
|
1017 |
||
| 52841 | 1018 |
primrec_new mirrror :: "('a, 'b, 'c) triple \<Rightarrow> ('c, 'b, 'a) triple" where
|
1019 |
"mirrror (Triple a b c) = Triple c b a" |
|
1020 |
||
| 53621 | 1021 |
text {*
|
1022 |
\noindent |
|
1023 |
The equations can be specified in any order, and it is acceptable to leave out |
|
1024 |
some cases, which are then unspecified. Pattern matching on the left-hand side |
|
1025 |
is restricted to a single datatype, which must correspond to the same argument |
|
1026 |
in all equations. |
|
1027 |
*} |
|
| 52828 | 1028 |
|
| 53621 | 1029 |
|
1030 |
subsubsection {* Simple Recursion
|
|
1031 |
\label{sssec:primrec-simple-recursion} *}
|
|
| 52828 | 1032 |
|
| 52841 | 1033 |
text {*
|
| 53621 | 1034 |
For simple recursive types, recursive calls on a constructor argument are |
1035 |
allowed on the right-hand side: |
|
| 52841 | 1036 |
*} |
1037 |
||
|
53330
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1038 |
primrec_new replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where |
|
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1039 |
"replicate Zero _ = []" | |
| 53644 | 1040 |
"replicate (Suc n) x = x # replicate n x" |
| 52841 | 1041 |
|
| 53621 | 1042 |
text {* \blankline *}
|
| 52843 | 1043 |
|
| 53332 | 1044 |
primrec_new at :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" where |
| 53644 | 1045 |
"at (x # xs) j = |
| 52843 | 1046 |
(case j of |
| 53644 | 1047 |
Zero \<Rightarrow> x |
1048 |
| Suc j' \<Rightarrow> at xs j')" |
|
| 52843 | 1049 |
|
| 53621 | 1050 |
text {* \blankline *}
|
1051 |
||
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1052 |
primrec_new (*<*)(in early) (*>*)tfold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) tlist \<Rightarrow> 'b" where
|
| 53644 | 1053 |
"tfold _ (TNil y) = y" | |
1054 |
"tfold f (TCons x xs) = f x (tfold f xs)" |
|
| 52841 | 1055 |
|
| 53025 | 1056 |
text {*
|
| 53621 | 1057 |
\noindent |
| 54402 | 1058 |
Pattern matching is only available for the argument on which the recursion takes |
1059 |
place. Fortunately, it is easy to generate pattern-maching equations using the |
|
1060 |
\keyw{simps\_of\_case} command provided by the theory
|
|
1061 |
\verb|~~/src/HOL/Library/Simps_Case_Conv|. |
|
1062 |
*} |
|
1063 |
||
1064 |
simps_of_case at_simps: at.simps |
|
1065 |
||
1066 |
text {*
|
|
1067 |
This generates the lemma collection @{thm [source] at_simps}:
|
|
1068 |
% |
|
1069 |
\[@{thm at_simps(1)[no_vars]}
|
|
1070 |
\qquad @{thm at_simps(2)[no_vars]}\]
|
|
1071 |
% |
|
| 54184 | 1072 |
The next example is defined using \keyw{fun} to escape the syntactic
|
1073 |
restrictions imposed on primitive recursive functions. The |
|
1074 |
@{command datatype_new_compat} command is needed to register new-style datatypes
|
|
1075 |
for use with \keyw{fun} and \keyw{function}
|
|
| 53621 | 1076 |
(Section~\ref{sssec:datatype-new-compat}):
|
| 53025 | 1077 |
*} |
| 52828 | 1078 |
|
| 53621 | 1079 |
datatype_new_compat nat |
1080 |
||
1081 |
text {* \blankline *}
|
|
1082 |
||
1083 |
fun at_least_two :: "nat \<Rightarrow> bool" where |
|
1084 |
"at_least_two (Suc (Suc _)) \<longleftrightarrow> True" | |
|
1085 |
"at_least_two _ \<longleftrightarrow> False" |
|
1086 |
||
1087 |
||
1088 |
subsubsection {* Mutual Recursion
|
|
1089 |
\label{sssec:primrec-mutual-recursion} *}
|
|
| 52828 | 1090 |
|
| 52841 | 1091 |
text {*
|
| 53621 | 1092 |
The syntax for mutually recursive functions over mutually recursive datatypes |
1093 |
is straightforward: |
|
| 52841 | 1094 |
*} |
1095 |
||
1096 |
primrec_new |
|
| 53623 | 1097 |
nat_of_even_nat :: "even_nat \<Rightarrow> nat" and |
1098 |
nat_of_odd_nat :: "odd_nat \<Rightarrow> nat" |
|
| 52841 | 1099 |
where |
| 53623 | 1100 |
"nat_of_even_nat Even_Zero = Zero" | |
1101 |
"nat_of_even_nat (Even_Suc n) = Suc (nat_of_odd_nat n)" | |
|
1102 |
"nat_of_odd_nat (Odd_Suc n) = Suc (nat_of_even_nat n)" |
|
| 52841 | 1103 |
|
| 53752 | 1104 |
text {* \blankline *}
|
1105 |
||
| 52841 | 1106 |
primrec_new |
|
53330
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1107 |
eval\<^sub>e :: "('a \<Rightarrow> int) \<Rightarrow> ('b \<Rightarrow> int) \<Rightarrow> ('a, 'b) exp \<Rightarrow> int" and
|
|
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1108 |
eval\<^sub>t :: "('a \<Rightarrow> int) \<Rightarrow> ('b \<Rightarrow> int) \<Rightarrow> ('a, 'b) trm \<Rightarrow> int" and
|
|
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1109 |
eval\<^sub>f :: "('a \<Rightarrow> int) \<Rightarrow> ('b \<Rightarrow> int) \<Rightarrow> ('a, 'b) fct \<Rightarrow> int"
|
| 52841 | 1110 |
where |
1111 |
"eval\<^sub>e \<gamma> \<xi> (Term t) = eval\<^sub>t \<gamma> \<xi> t" | |
|
1112 |
"eval\<^sub>e \<gamma> \<xi> (Sum t e) = eval\<^sub>t \<gamma> \<xi> t + eval\<^sub>e \<gamma> \<xi> e" | |
|
|
53330
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1113 |
"eval\<^sub>t \<gamma> \<xi> (Factor f) = eval\<^sub>f \<gamma> \<xi> f" | |
| 52841 | 1114 |
"eval\<^sub>t \<gamma> \<xi> (Prod f t) = eval\<^sub>f \<gamma> \<xi> f + eval\<^sub>t \<gamma> \<xi> t" | |
1115 |
"eval\<^sub>f \<gamma> _ (Const a) = \<gamma> a" | |
|
1116 |
"eval\<^sub>f _ \<xi> (Var b) = \<xi> b" | |
|
1117 |
"eval\<^sub>f \<gamma> \<xi> (Expr e) = eval\<^sub>e \<gamma> \<xi> e" |
|
1118 |
||
| 53621 | 1119 |
text {*
|
1120 |
\noindent |
|
| 53647 | 1121 |
Mutual recursion is possible within a single type, using \keyw{fun}:
|
| 53621 | 1122 |
*} |
| 52828 | 1123 |
|
| 53621 | 1124 |
fun |
1125 |
even :: "nat \<Rightarrow> bool" and |
|
1126 |
odd :: "nat \<Rightarrow> bool" |
|
1127 |
where |
|
1128 |
"even Zero = True" | |
|
1129 |
"even (Suc n) = odd n" | |
|
1130 |
"odd Zero = False" | |
|
1131 |
"odd (Suc n) = even n" |
|
1132 |
||
1133 |
||
1134 |
subsubsection {* Nested Recursion
|
|
1135 |
\label{sssec:primrec-nested-recursion} *}
|
|
1136 |
||
1137 |
text {*
|
|
1138 |
In a departure from the old datatype package, nested recursion is normally |
|
1139 |
handled via the map functions of the nesting type constructors. For example, |
|
1140 |
recursive calls are lifted to lists using @{const map}:
|
|
1141 |
*} |
|
| 52828 | 1142 |
|
| 52843 | 1143 |
(*<*) |
| 53644 | 1144 |
datatype_new 'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f (lbl\<^sub>f\<^sub>f: 'a) (sub\<^sub>f\<^sub>f: "'a tree\<^sub>f\<^sub>f list") |
| 52843 | 1145 |
(*>*) |
| 53028 | 1146 |
primrec_new at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \<Rightarrow> nat list \<Rightarrow> 'a" where |
1147 |
"at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js = |
|
| 52843 | 1148 |
(case js of |
1149 |
[] \<Rightarrow> a |
|
| 53028 | 1150 |
| j # js' \<Rightarrow> at (map (\<lambda>t. at\<^sub>f\<^sub>f t js') ts) j)" |
| 52843 | 1151 |
|
| 53025 | 1152 |
text {*
|
| 53647 | 1153 |
\noindent |
| 53621 | 1154 |
The next example features recursion through the @{text option} type. Although
|
| 53623 | 1155 |
@{text option} is not a new-style datatype, it is registered as a BNF with the
|
| 54491 | 1156 |
map function @{const map_option}:
|
| 53025 | 1157 |
*} |
| 52843 | 1158 |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1159 |
primrec_new (*<*)(in early) (*>*)sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" where
|
| 52843 | 1160 |
"sum_btree (BNode a lt rt) = |
| 54491 | 1161 |
a + the_default 0 (map_option sum_btree lt) + |
1162 |
the_default 0 (map_option sum_btree rt)" |
|
| 52843 | 1163 |
|
| 53136 | 1164 |
text {*
|
| 53621 | 1165 |
\noindent |
1166 |
The same principle applies for arbitrary type constructors through which |
|
1167 |
recursion is possible. Notably, the map function for the function type |
|
1168 |
(@{text \<Rightarrow>}) is simply composition (@{text "op \<circ>"}):
|
|
| 53136 | 1169 |
*} |
| 52828 | 1170 |
|
| 54182 | 1171 |
primrec_new (*<*)(in early) (*>*)relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
|
1172 |
"relabel_ft f (FTLeaf x) = FTLeaf (f x)" | |
|
1173 |
"relabel_ft f (FTNode g) = FTNode (relabel_ft f \<circ> g)" |
|
1174 |
||
1175 |
text {*
|
|
1176 |
\noindent |
|
1177 |
For convenience, recursion through functions can also be expressed using |
|
1178 |
$\lambda$-abstractions and function application rather than through composition. |
|
1179 |
For example: |
|
1180 |
*} |
|
1181 |
||
1182 |
primrec_new relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
|
|
1183 |
"relabel_ft f (FTLeaf x) = FTLeaf (f x)" | |
|
1184 |
"relabel_ft f (FTNode g) = FTNode (\<lambda>x. relabel_ft f (g x))" |
|
| 52828 | 1185 |
|
| 54183 | 1186 |
text {* \blankline *}
|
1187 |
||
1188 |
primrec_new subtree_ft :: "'a \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where |
|
1189 |
"subtree_ft x (FTNode g) = g x" |
|
1190 |
||
| 52843 | 1191 |
text {*
|
| 53621 | 1192 |
\noindent |
| 54182 | 1193 |
For recursion through curried $n$-ary functions, $n$ applications of |
1194 |
@{term "op \<circ>"} are necessary. The examples below illustrate the case where
|
|
1195 |
$n = 2$: |
|
| 53621 | 1196 |
*} |
1197 |
||
| 54182 | 1198 |
datatype_new 'a ftree2 = FTLeaf2 'a | FTNode2 "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2" |
1199 |
||
1200 |
text {* \blankline *}
|
|
1201 |
||
1202 |
primrec_new (*<*)(in early) (*>*)relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
|
|
1203 |
"relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" | |
|
1204 |
"relabel_ft2 f (FTNode2 g) = FTNode2 (op \<circ> (op \<circ> (relabel_ft2 f)) g)" |
|
1205 |
||
1206 |
text {* \blankline *}
|
|
1207 |
||
1208 |
primrec_new relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
|
|
1209 |
"relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" | |
|
1210 |
"relabel_ft2 f (FTNode2 g) = FTNode2 (\<lambda>x y. relabel_ft2 f (g x y))" |
|
| 54031 | 1211 |
|
| 54183 | 1212 |
text {* \blankline *}
|
1213 |
||
1214 |
primrec_new subtree_ft2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where |
|
1215 |
"subtree_ft2 x y (FTNode2 g) = g x y" |
|
1216 |
||
| 53621 | 1217 |
|
1218 |
subsubsection {* Nested-as-Mutual Recursion
|
|
| 53644 | 1219 |
\label{sssec:primrec-nested-as-mutual-recursion} *}
|
| 53621 | 1220 |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1221 |
(*<*) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1222 |
locale n2m begin |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1223 |
(*>*) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1224 |
|
| 53621 | 1225 |
text {*
|
1226 |
For compatibility with the old package, but also because it is sometimes |
|
1227 |
convenient in its own right, it is possible to treat nested recursive datatypes |
|
1228 |
as mutually recursive ones if the recursion takes place though new-style |
|
1229 |
datatypes. For example: |
|
| 52843 | 1230 |
*} |
1231 |
||
|
53331
20440c789759
prove theorem in the right context (that knows about local variables)
traytel
parents:
53330
diff
changeset
|
1232 |
primrec_new |
| 53647 | 1233 |
at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \<Rightarrow> nat list \<Rightarrow> 'a" and |
1234 |
ats\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f list \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> 'a" |
|
| 52843 | 1235 |
where |
| 53647 | 1236 |
"at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js = |
| 52843 | 1237 |
(case js of |
1238 |
[] \<Rightarrow> a |
|
| 53647 | 1239 |
| j # js' \<Rightarrow> ats\<^sub>f\<^sub>f ts j js')" | |
1240 |
"ats\<^sub>f\<^sub>f (t # ts) j = |
|
| 52843 | 1241 |
(case j of |
| 53647 | 1242 |
Zero \<Rightarrow> at\<^sub>f\<^sub>f t |
1243 |
| Suc j' \<Rightarrow> ats\<^sub>f\<^sub>f ts j')" |
|
| 52843 | 1244 |
|
| 53647 | 1245 |
text {*
|
1246 |
\noindent |
|
| 54287 | 1247 |
Appropriate induction rules are generated as |
| 54031 | 1248 |
@{thm [source] at\<^sub>f\<^sub>f.induct},
|
1249 |
@{thm [source] ats\<^sub>f\<^sub>f.induct}, and
|
|
| 54287 | 1250 |
@{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The
|
1251 |
induction rules and the underlying recursors are generated on a per-need basis |
|
1252 |
and are kept in a cache to speed up subsequent definitions. |
|
| 53647 | 1253 |
|
1254 |
Here is a second example: |
|
1255 |
*} |
|
| 53621 | 1256 |
|
|
53331
20440c789759
prove theorem in the right context (that knows about local variables)
traytel
parents:
53330
diff
changeset
|
1257 |
primrec_new |
|
53330
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1258 |
sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" and
|
|
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1259 |
sum_btree_option :: "'a btree option \<Rightarrow> 'a" |
| 52843 | 1260 |
where |
1261 |
"sum_btree (BNode a lt rt) = |
|
| 53025 | 1262 |
a + sum_btree_option lt + sum_btree_option rt" | |
|
53330
77da8d3c46e0
fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents:
53262
diff
changeset
|
1263 |
"sum_btree_option None = 0" | |
| 53025 | 1264 |
"sum_btree_option (Some t) = sum_btree t" |
| 52843 | 1265 |
|
1266 |
text {*
|
|
| 53621 | 1267 |
% * can pretend a nested type is mutually recursive (if purely inductive) |
1268 |
% * avoids the higher-order map |
|
1269 |
% * e.g. |
|
1270 |
||
| 53617 | 1271 |
% * this can always be avoided; |
1272 |
% * e.g. in our previous example, we first mapped the recursive |
|
1273 |
% calls, then we used a generic at function to retrieve the result |
|
1274 |
% |
|
1275 |
% * there's no hard-and-fast rule of when to use one or the other, |
|
1276 |
% just like there's no rule when to use fold and when to use |
|
1277 |
% primrec_new |
|
1278 |
% |
|
1279 |
% * higher-order approach, considering nesting as nesting, is more |
|
1280 |
% compositional -- e.g. we saw how we could reuse an existing polymorphic |
|
| 53647 | 1281 |
% at or the_default, whereas @{const ats\<^sub>f\<^sub>f} is much more specific
|
| 53617 | 1282 |
% |
1283 |
% * but: |
|
1284 |
% * is perhaps less intuitive, because it requires higher-order thinking |
|
1285 |
% * may seem inefficient, and indeed with the code generator the |
|
1286 |
% mutually recursive version might be nicer |
|
1287 |
% * is somewhat indirect -- must apply a map first, then compute a result |
|
1288 |
% (cannot mix) |
|
| 53647 | 1289 |
% * the auxiliary functions like @{const ats\<^sub>f\<^sub>f} are sometimes useful in own right
|
| 53617 | 1290 |
% |
1291 |
% * impact on automation unclear |
|
1292 |
% |
|
| 52843 | 1293 |
*} |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1294 |
(*<*) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1295 |
end |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1296 |
(*>*) |
| 52843 | 1297 |
|
| 52824 | 1298 |
|
| 53617 | 1299 |
subsection {* Command Syntax
|
1300 |
\label{ssec:primrec-command-syntax} *}
|
|
1301 |
||
1302 |
||
| 53621 | 1303 |
subsubsection {* \keyw{primrec\_new}
|
1304 |
\label{sssec:primrec-new} *}
|
|
| 52828 | 1305 |
|
1306 |
text {*
|
|
| 53829 | 1307 |
\begin{matharray}{rcl}
|
1308 |
@{command_def "primrec_new"} & : & @{text "local_theory \<rightarrow> local_theory"}
|
|
1309 |
\end{matharray}
|
|
| 52794 | 1310 |
|
| 52840 | 1311 |
@{rail "
|
| 53829 | 1312 |
@@{command primrec_new} target? fixes \\ @'where' (@{syntax pr_equation} + '|')
|
| 52840 | 1313 |
; |
| 53829 | 1314 |
@{syntax_def pr_equation}: thmdecl? prop
|
| 52840 | 1315 |
"} |
| 52828 | 1316 |
*} |
1317 |
||
| 52840 | 1318 |
|
| 53619 | 1319 |
(* |
| 52840 | 1320 |
subsection {* Generated Theorems
|
1321 |
\label{ssec:primrec-generated-theorems} *}
|
|
| 52824 | 1322 |
|
| 52828 | 1323 |
text {*
|
| 53617 | 1324 |
% * synthesized nonrecursive definition |
1325 |
% * user specification is rederived from it, exactly as entered |
|
1326 |
% |
|
1327 |
% * induct |
|
1328 |
% * mutualized |
|
1329 |
% * without some needless induction hypotheses if not used |
|
1330 |
% * fold, rec |
|
1331 |
% * mutualized |
|
| 52828 | 1332 |
*} |
| 53619 | 1333 |
*) |
1334 |
||
| 52824 | 1335 |
|
| 52840 | 1336 |
subsection {* Recursive Default Values for Selectors
|
| 53623 | 1337 |
\label{ssec:primrec-recursive-default-values-for-selectors} *}
|
| 52827 | 1338 |
|
1339 |
text {*
|
|
1340 |
A datatype selector @{text un_D} can have a default value for each constructor
|
|
1341 |
on which it is not otherwise specified. Occasionally, it is useful to have the |
|
1342 |
default value be defined recursively. This produces a chicken-and-egg situation |
|
| 53621 | 1343 |
that may seem unsolvable, because the datatype is not introduced yet at the |
| 52827 | 1344 |
moment when the selectors are introduced. Of course, we can always define the |
1345 |
selectors manually afterward, but we then have to state and prove all the |
|
1346 |
characteristic theorems ourselves instead of letting the package do it. |
|
1347 |
||
1348 |
Fortunately, there is a fairly elegant workaround that relies on overloading and |
|
1349 |
that avoids the tedium of manual derivations: |
|
1350 |
||
1351 |
\begin{enumerate}
|
|
1352 |
\setlength{\itemsep}{0pt}
|
|
1353 |
||
1354 |
\item |
|
1355 |
Introduce a fully unspecified constant @{text "un_D\<^sub>0 \<Colon> 'a"} using
|
|
1356 |
@{keyword consts}.
|
|
1357 |
||
1358 |
\item |
|
| 53535 | 1359 |
Define the datatype, specifying @{text "un_D\<^sub>0"} as the selector's default
|
1360 |
value. |
|
| 52827 | 1361 |
|
1362 |
\item |
|
| 53535 | 1363 |
Define the behavior of @{text "un_D\<^sub>0"} on values of the newly introduced
|
1364 |
datatype using the \keyw{overloading} command.
|
|
| 52827 | 1365 |
|
1366 |
\item |
|
1367 |
Derive the desired equation on @{text un_D} from the characteristic equations
|
|
1368 |
for @{text "un_D\<^sub>0"}.
|
|
1369 |
\end{enumerate}
|
|
1370 |
||
| 53619 | 1371 |
\noindent |
| 52827 | 1372 |
The following example illustrates this procedure: |
1373 |
*} |
|
1374 |
||
1375 |
consts termi\<^sub>0 :: 'a |
|
1376 |
||
| 53619 | 1377 |
text {* \blankline *}
|
1378 |
||
| 53491 | 1379 |
datatype_new ('a, 'b) tlist =
|
| 52827 | 1380 |
TNil (termi: 'b) (defaults ttl: TNil) |
| 53491 | 1381 |
| TCons (thd: 'a) (ttl : "('a, 'b) tlist") (defaults termi: "\<lambda>_ xs. termi\<^sub>0 xs")
|
| 52827 | 1382 |
|
| 53619 | 1383 |
text {* \blankline *}
|
1384 |
||
| 52827 | 1385 |
overloading |
| 53491 | 1386 |
termi\<^sub>0 \<equiv> "termi\<^sub>0 \<Colon> ('a, 'b) tlist \<Rightarrow> 'b"
|
| 52827 | 1387 |
begin |
| 53491 | 1388 |
primrec_new termi\<^sub>0 :: "('a, 'b) tlist \<Rightarrow> 'b" where
|
| 53621 | 1389 |
"termi\<^sub>0 (TNil y) = y" | |
1390 |
"termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs" |
|
| 52827 | 1391 |
end |
1392 |
||
| 53619 | 1393 |
text {* \blankline *}
|
1394 |
||
| 52827 | 1395 |
lemma terminal_TCons[simp]: "termi (TCons x xs) = termi xs" |
1396 |
by (cases xs) auto |
|
1397 |
||
1398 |
||
| 52828 | 1399 |
subsection {* Compatibility Issues
|
| 53617 | 1400 |
\label{ssec:primrec-compatibility-issues} *}
|
| 52828 | 1401 |
|
1402 |
text {*
|
|
| 53997 | 1403 |
The command @{command primrec_new} has been designed to be highly compatible
|
1404 |
with the old \keyw{primrec}, to ease migration. There is nonetheless at least
|
|
1405 |
one incompatibility that may arise when porting to the new package: |
|
1406 |
||
1407 |
\begin{itemize}
|
|
1408 |
\setlength{\itemsep}{0pt}
|
|
1409 |
||
| 54185 | 1410 |
\item \emph{Some theorems have different names.}
|
| 53997 | 1411 |
For $m > 1$ mutually recursive functions, |
|
54023
cede3c1d2417
minor doc fix (there is no guarantee that the equations for a given f_i are contiguous in the collection)
blanchet
parents:
54014
diff
changeset
|
1412 |
@{text "f\<^sub>1_\<dots>_f\<^sub>m.simps"} has been broken down into separate
|
|
cede3c1d2417
minor doc fix (there is no guarantee that the equations for a given f_i are contiguous in the collection)
blanchet
parents:
54014
diff
changeset
|
1413 |
subcollections @{text "f\<^sub>i.simps"}.
|
| 53997 | 1414 |
\end{itemize}
|
| 52828 | 1415 |
*} |
| 52794 | 1416 |
|
1417 |
||
| 52827 | 1418 |
section {* Defining Codatatypes
|
| 52805 | 1419 |
\label{sec:defining-codatatypes} *}
|
1420 |
||
1421 |
text {*
|
|
| 53829 | 1422 |
Codatatypes can be specified using the @{command codatatype} command. The
|
| 53623 | 1423 |
command is first illustrated through concrete examples featuring different |
1424 |
flavors of corecursion. More examples can be found in the directory |
|
| 53997 | 1425 |
\verb|~~/src/HOL/|\allowbreak\verb|BNF/Examples|. The |
1426 |
\emph{Archive of Formal Proofs} also includes some useful codatatypes, notably
|
|
1427 |
for lazy lists \cite{lochbihler-2010}.
|
|
| 52805 | 1428 |
*} |
| 52792 | 1429 |
|
| 52824 | 1430 |
|
| 53617 | 1431 |
subsection {* Introductory Examples
|
1432 |
\label{ssec:codatatype-introductory-examples} *}
|
|
| 52794 | 1433 |
|
| 53623 | 1434 |
|
1435 |
subsubsection {* Simple Corecursion
|
|
1436 |
\label{sssec:codatatype-simple-corecursion} *}
|
|
1437 |
||
| 52805 | 1438 |
text {*
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
1439 |
Noncorecursive codatatypes coincide with the corresponding datatypes, so they |
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
1440 |
are useless in practice. \emph{Corecursive codatatypes} have the same syntax
|
| 53623 | 1441 |
as recursive datatypes, except for the command name. For example, here is the |
1442 |
definition of lazy lists: |
|
1443 |
*} |
|
1444 |
||
1445 |
codatatype (lset: 'a) llist (map: lmap rel: llist_all2) = |
|
1446 |
lnull: LNil (defaults ltl: LNil) |
|
1447 |
| LCons (lhd: 'a) (ltl: "'a llist") |
|
1448 |
||
1449 |
text {*
|
|
1450 |
\noindent |
|
1451 |
Lazy lists can be infinite, such as @{text "LCons 0 (LCons 0 (\<dots>))"} and
|
|
| 53647 | 1452 |
@{text "LCons 0 (LCons 1 (LCons 2 (\<dots>)))"}. Here is a related type, that of
|
1453 |
infinite streams: |
|
1454 |
*} |
|
1455 |
||
1456 |
codatatype (sset: 'a) stream (map: smap rel: stream_all2) = |
|
1457 |
SCons (shd: 'a) (stl: "'a stream") |
|
1458 |
||
1459 |
text {*
|
|
1460 |
\noindent |
|
1461 |
Another interesting type that can |
|
| 53623 | 1462 |
be defined as a codatatype is that of the extended natural numbers: |
1463 |
*} |
|
1464 |
||
| 53644 | 1465 |
codatatype enat = EZero | ESuc enat |
| 53623 | 1466 |
|
1467 |
text {*
|
|
1468 |
\noindent |
|
1469 |
This type has exactly one infinite element, @{text "ESuc (ESuc (ESuc (\<dots>)))"},
|
|
1470 |
that represents $\infty$. In addition, it has finite values of the form |
|
1471 |
@{text "ESuc (\<dots> (ESuc EZero)\<dots>)"}.
|
|
| 53675 | 1472 |
|
1473 |
Here is an example with many constructors: |
|
| 52805 | 1474 |
*} |
| 53623 | 1475 |
|
| 53675 | 1476 |
codatatype 'a process = |
1477 |
Fail |
|
1478 |
| Skip (cont: "'a process") |
|
1479 |
| Action (prefix: 'a) (cont: "'a process") |
|
1480 |
| Choice (left: "'a process") (right: "'a process") |
|
1481 |
||
| 53750 | 1482 |
text {*
|
| 53829 | 1483 |
\noindent |
| 53750 | 1484 |
Notice that the @{const cont} selector is associated with both @{const Skip}
|
| 54146 | 1485 |
and @{const Action}.
|
| 53750 | 1486 |
*} |
1487 |
||
| 53623 | 1488 |
|
1489 |
subsubsection {* Mutual Corecursion
|
|
1490 |
\label{sssec:codatatype-mutual-corecursion} *}
|
|
1491 |
||
1492 |
text {*
|
|
1493 |
\noindent |
|
1494 |
The example below introduces a pair of \emph{mutually corecursive} types:
|
|
1495 |
*} |
|
1496 |
||
1497 |
codatatype even_enat = Even_EZero | Even_ESuc odd_enat |
|
1498 |
and odd_enat = Odd_ESuc even_enat |
|
1499 |
||
1500 |
||
1501 |
subsubsection {* Nested Corecursion
|
|
1502 |
\label{sssec:codatatype-nested-corecursion} *}
|
|
1503 |
||
1504 |
text {*
|
|
1505 |
\noindent |
|
| 53675 | 1506 |
The next examples feature \emph{nested corecursion}:
|
| 53623 | 1507 |
*} |
1508 |
||
| 53644 | 1509 |
codatatype 'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i (lbl\<^sub>i\<^sub>i: 'a) (sub\<^sub>i\<^sub>i: "'a tree\<^sub>i\<^sub>i llist") |
| 53675 | 1510 |
|
| 53752 | 1511 |
text {* \blankline *}
|
1512 |
||
| 53644 | 1513 |
codatatype 'a tree\<^sub>i\<^sub>s = Node\<^sub>i\<^sub>s (lbl\<^sub>i\<^sub>s: 'a) (sub\<^sub>i\<^sub>s: "'a tree\<^sub>i\<^sub>s fset") |
| 52805 | 1514 |
|
| 53752 | 1515 |
text {* \blankline *}
|
1516 |
||
| 53675 | 1517 |
codatatype 'a state_machine = |
| 53751 | 1518 |
State_Machine (accept: bool) (trans: "'a \<Rightarrow> 'a state_machine") |
| 53675 | 1519 |
|
| 52824 | 1520 |
|
| 53617 | 1521 |
subsection {* Command Syntax
|
1522 |
\label{ssec:codatatype-command-syntax} *}
|
|
| 52805 | 1523 |
|
| 53619 | 1524 |
|
| 53621 | 1525 |
subsubsection {* \keyw{codatatype}
|
1526 |
\label{sssec:codatatype} *}
|
|
| 53619 | 1527 |
|
| 52824 | 1528 |
text {*
|
| 53829 | 1529 |
\begin{matharray}{rcl}
|
1530 |
@{command_def "codatatype"} & : & @{text "local_theory \<rightarrow> local_theory"}
|
|
1531 |
\end{matharray}
|
|
1532 |
||
1533 |
@{rail "
|
|
1534 |
@@{command codatatype} target? \\
|
|
1535 |
(@{syntax dt_name} '=' (@{syntax ctor} + '|') + @'and')
|
|
1536 |
"} |
|
1537 |
||
1538 |
\noindent |
|
| 52827 | 1539 |
Definitions of codatatypes have almost exactly the same syntax as for datatypes |
| 53829 | 1540 |
(Section~\ref{ssec:datatype-command-syntax}). The @{text "no_discs_sels"} option
|
1541 |
is not available, because destructors are a crucial notion for codatatypes. |
|
| 53623 | 1542 |
*} |
1543 |
||
1544 |
||
1545 |
subsection {* Generated Constants
|
|
1546 |
\label{ssec:codatatype-generated-constants} *}
|
|
1547 |
||
1548 |
text {*
|
|
1549 |
Given a codatatype @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
|
|
1550 |
with $m > 0$ live type variables and $n$ constructors @{text "t.C\<^sub>1"},
|
|
1551 |
\ldots, @{text "t.C\<^sub>n"}, the same auxiliary constants are generated as for
|
|
1552 |
datatypes (Section~\ref{ssec:datatype-generated-constants}), except that the
|
|
1553 |
iterator and the recursor are replaced by dual concepts: |
|
1554 |
||
1555 |
\begin{itemize}
|
|
1556 |
\setlength{\itemsep}{0pt}
|
|
1557 |
||
| 54494 | 1558 |
\item \relax{Coiterator}: @{text unfold_t}
|
1559 |
||
1560 |
\item \relax{Corecursor}: @{text corec_t}
|
|
| 53623 | 1561 |
|
1562 |
\end{itemize}
|
|
1563 |
*} |
|
1564 |
||
1565 |
||
1566 |
subsection {* Generated Theorems
|
|
1567 |
\label{ssec:codatatype-generated-theorems} *}
|
|
1568 |
||
1569 |
text {*
|
|
| 53829 | 1570 |
The characteristic theorems generated by @{command codatatype} are grouped in
|
| 53623 | 1571 |
three broad categories: |
1572 |
||
1573 |
\begin{itemize}
|
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1574 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1575 |
|
| 53623 | 1576 |
\item The \emph{free constructor theorems} are properties about the constructors
|
1577 |
and destructors that can be derived for any freely generated type. |
|
1578 |
||
1579 |
\item The \emph{functorial theorems} are properties of datatypes related to
|
|
1580 |
their BNF nature. |
|
1581 |
||
1582 |
\item The \emph{coinductive theorems} are properties of datatypes related to
|
|
1583 |
their coinductive nature. |
|
1584 |
\end{itemize}
|
|
1585 |
||
1586 |
\noindent |
|
1587 |
The first two categories are exactly as for datatypes and are described in |
|
| 53642 | 1588 |
Sections |
1589 |
\ref{sssec:free-constructor-theorems}~and~\ref{sssec:functorial-theorems}.
|
|
| 52824 | 1590 |
*} |
1591 |
||
| 53617 | 1592 |
|
| 53623 | 1593 |
subsubsection {* Coinductive Theorems
|
1594 |
\label{sssec:coinductive-theorems} *}
|
|
1595 |
||
1596 |
text {*
|
|
| 54031 | 1597 |
The coinductive theorems are listed below for @{typ "'a llist"}:
|
| 53623 | 1598 |
|
1599 |
\begin{indentblock}
|
|
1600 |
\begin{description}
|
|
1601 |
||
| 53643 | 1602 |
\item[\begin{tabular}{@ {}l@ {}}
|
1603 |
@{text "t."}\hthm{coinduct} @{text "[coinduct t, consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
|
|
1604 |
\phantom{@{text "t."}\hthm{coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm:
|
|
1605 |
\end{tabular}] ~ \\
|
|
| 53623 | 1606 |
@{thm llist.coinduct[no_vars]}
|
| 53617 | 1607 |
|
| 53643 | 1608 |
\item[\begin{tabular}{@ {}l@ {}}
|
1609 |
@{text "t."}\hthm{strong\_coinduct} @{text "[consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
|
|
1610 |
\phantom{@{text "t."}\hthm{strong\_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm:
|
|
1611 |
\end{tabular}] ~ \\
|
|
1612 |
@{thm llist.strong_coinduct[no_vars]}
|
|
| 53617 | 1613 |
|
| 53643 | 1614 |
\item[\begin{tabular}{@ {}l@ {}}
|
1615 |
@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m, case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"} \\
|
|
1616 |
@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{strong\_coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
|
|
1617 |
\phantom{@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{strong\_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm:
|
|
1618 |
\end{tabular}] ~ \\
|
|
1619 |
Given $m > 1$ mutually corecursive codatatypes, these coinduction rules can be |
|
1620 |
used to prove $m$ properties simultaneously. |
|
1621 |
||
| 54031 | 1622 |
\item[@{text "t."}\hthm{unfold}\rm:] ~ \\
|
| 53623 | 1623 |
@{thm llist.unfold(1)[no_vars]} \\
|
1624 |
@{thm llist.unfold(2)[no_vars]}
|
|
1625 |
||
| 54031 | 1626 |
\item[@{text "t."}\hthm{corec}\rm:] ~ \\
|
| 53623 | 1627 |
@{thm llist.corec(1)[no_vars]} \\
|
1628 |
@{thm llist.corec(2)[no_vars]}
|
|
1629 |
||
| 53703 | 1630 |
\item[@{text "t."}\hthm{disc\_unfold}\rm:] ~ \\
|
| 53643 | 1631 |
@{thm llist.disc_unfold(1)[no_vars]} \\
|
1632 |
@{thm llist.disc_unfold(2)[no_vars]}
|
|
1633 |
||
| 53703 | 1634 |
\item[@{text "t."}\hthm{disc\_corec}\rm:] ~ \\
|
| 53643 | 1635 |
@{thm llist.disc_corec(1)[no_vars]} \\
|
1636 |
@{thm llist.disc_corec(2)[no_vars]}
|
|
1637 |
||
1638 |
\item[@{text "t."}\hthm{disc\_unfold\_iff} @{text "[simp]"}\rm:] ~ \\
|
|
1639 |
@{thm llist.disc_unfold_iff(1)[no_vars]} \\
|
|
1640 |
@{thm llist.disc_unfold_iff(2)[no_vars]}
|
|
1641 |
||
1642 |
\item[@{text "t."}\hthm{disc\_corec\_iff} @{text "[simp]"}\rm:] ~ \\
|
|
1643 |
@{thm llist.disc_corec_iff(1)[no_vars]} \\
|
|
1644 |
@{thm llist.disc_corec_iff(2)[no_vars]}
|
|
1645 |
||
1646 |
\item[@{text "t."}\hthm{sel\_unfold} @{text "[simp]"}\rm:] ~ \\
|
|
1647 |
@{thm llist.sel_unfold(1)[no_vars]} \\
|
|
1648 |
@{thm llist.sel_unfold(2)[no_vars]}
|
|
1649 |
||
1650 |
\item[@{text "t."}\hthm{sel\_corec} @{text "[simp]"}\rm:] ~ \\
|
|
1651 |
@{thm llist.sel_corec(1)[no_vars]} \\
|
|
1652 |
@{thm llist.sel_corec(2)[no_vars]}
|
|
1653 |
||
| 53623 | 1654 |
\end{description}
|
1655 |
\end{indentblock}
|
|
1656 |
||
1657 |
\noindent |
|
| 53829 | 1658 |
For convenience, @{command codatatype} also provides the following collection:
|
| 53623 | 1659 |
|
1660 |
\begin{indentblock}
|
|
1661 |
\begin{description}
|
|
1662 |
||
| 54031 | 1663 |
\item[@{text "t."}\hthm{simps} = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.disc_corec} @{text t.disc_corec_iff}] ~ \\
|
1664 |
@{text t.sel_corec} @{text t.disc_unfold} @{text t.disc_unfold_iff} @{text t.sel_unfold} @{text t.map} \\
|
|
1665 |
@{text t.rel_inject} @{text t.rel_distinct} @{text t.set}
|
|
| 53623 | 1666 |
|
1667 |
\end{description}
|
|
1668 |
\end{indentblock}
|
|
1669 |
*} |
|
| 52805 | 1670 |
|
1671 |
||
| 52827 | 1672 |
section {* Defining Corecursive Functions
|
| 52805 | 1673 |
\label{sec:defining-corecursive-functions} *}
|
1674 |
||
1675 |
text {*
|
|
| 54183 | 1676 |
Corecursive functions can be specified using the @{command primcorec} and
|
1677 |
\keyw{prim\-corec\-ursive} commands, which support primitive corecursion, or
|
|
1678 |
using the more general \keyw{partial\_function} command. Here, the focus is on
|
|
1679 |
the first two. More examples can be found in the directory |
|
|
53753
ae7f50e70c09
renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents:
53752
diff
changeset
|
1680 |
\verb|~~/src/HOL/BNF/Examples|. |
| 53644 | 1681 |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1682 |
Whereas recursive functions consume datatypes one constructor at a time, |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1683 |
corecursive functions construct codatatypes one constructor at a time. |
| 53752 | 1684 |
Partly reflecting a lack of agreement among proponents of coalgebraic methods, |
1685 |
Isabelle supports three competing syntaxes for specifying a function $f$: |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1686 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1687 |
\begin{itemize}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1688 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1689 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1690 |
\abovedisplayskip=.5\abovedisplayskip |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1691 |
\belowdisplayskip=.5\belowdisplayskip |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1692 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1693 |
\item The \emph{destructor view} specifies $f$ by implications of the form
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1694 |
\[@{text "\<dots> \<Longrightarrow> is_C\<^sub>j (f x\<^sub>1 \<dots> x\<^sub>n)"}\] and
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1695 |
equations of the form |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1696 |
\[@{text "un_C\<^sub>ji (f x\<^sub>1 \<dots> x\<^sub>n) = \<dots>"}\]
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1697 |
This style is popular in the coalgebraic literature. |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1698 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1699 |
\item The \emph{constructor view} specifies $f$ by equations of the form
|
| 54183 | 1700 |
\[@{text "\<dots> \<Longrightarrow> f x\<^sub>1 \<dots> x\<^sub>n = C\<^sub>j \<dots>"}\]
|
| 53752 | 1701 |
This style is often more concise than the previous one. |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1702 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1703 |
\item The \emph{code view} specifies $f$ by a single equation of the form
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1704 |
\[@{text "f x\<^sub>1 \<dots> x\<^sub>n = \<dots>"}\]
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1705 |
with restrictions on the format of the right-hand side. Lazy functional |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1706 |
programming languages such as Haskell support a generalized version of this |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1707 |
style. |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1708 |
\end{itemize}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1709 |
|
|
53753
ae7f50e70c09
renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents:
53752
diff
changeset
|
1710 |
All three styles are available as input syntax. Whichever syntax is chosen, |
|
ae7f50e70c09
renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents:
53752
diff
changeset
|
1711 |
characteristic theorems for all three styles are generated. |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1712 |
|
| 52828 | 1713 |
%%% TODO: partial_function? E.g. for defining tail recursive function on lazy |
1714 |
%%% lists (cf. terminal0 in TLList.thy) |
|
| 52805 | 1715 |
*} |
1716 |
||
| 52824 | 1717 |
|
| 53617 | 1718 |
subsection {* Introductory Examples
|
1719 |
\label{ssec:primcorec-introductory-examples} *}
|
|
| 52805 | 1720 |
|
| 53646 | 1721 |
text {*
|
1722 |
Primitive corecursion is illustrated through concrete examples based on the |
|
1723 |
codatatypes defined in Section~\ref{ssec:codatatype-introductory-examples}. More
|
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1724 |
examples can be found in the directory \verb|~~/src/HOL/BNF/Examples|. The code |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1725 |
view is favored in the examples below. Sections |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1726 |
\ref{ssec:primrec-constructor-view} and \ref{ssec:primrec-destructor-view}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1727 |
present the same examples expressed using the constructor and destructor views. |
| 53646 | 1728 |
*} |
1729 |
||
| 53644 | 1730 |
subsubsection {* Simple Corecursion
|
1731 |
\label{sssec:primcorec-simple-corecursion} *}
|
|
1732 |
||
| 53646 | 1733 |
text {*
|
| 53752 | 1734 |
Following the code view, corecursive calls are allowed on the right-hand side as |
1735 |
long as they occur under a constructor, which itself appears either directly to |
|
1736 |
the right of the equal sign or in a conditional expression: |
|
| 53646 | 1737 |
*} |
1738 |
||
| 53826 | 1739 |
primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
|
| 54072 | 1740 |
"literate g x = LCons x (literate g (g x))" |
| 53647 | 1741 |
|
| 53677 | 1742 |
text {* \blankline *}
|
1743 |
||
| 53826 | 1744 |
primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
|
| 54072 | 1745 |
"siterate g x = SCons x (siterate g (g x))" |
| 53644 | 1746 |
|
| 53646 | 1747 |
text {*
|
1748 |
\noindent |
|
1749 |
The constructor ensures that progress is made---i.e., the function is |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1750 |
\emph{productive}. The above functions compute the infinite lazy list or stream
|
| 54072 | 1751 |
@{text "[x, g x, g (g x), \<dots>]"}. Productivity guarantees that prefixes
|
1752 |
@{text "[x, g x, g (g x), \<dots>, (g ^^ k) x]"} of arbitrary finite length
|
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1753 |
@{text k} can be computed by unfolding the code equation a finite number of
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
1754 |
times. |
| 53646 | 1755 |
|
| 53752 | 1756 |
Corecursive functions construct codatatype values, but nothing prevents them |
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
1757 |
from also consuming such values. The following function drops every second |
| 53675 | 1758 |
element in a stream: |
1759 |
*} |
|
1760 |
||
| 53826 | 1761 |
primcorec every_snd :: "'a stream \<Rightarrow> 'a stream" where |
| 53675 | 1762 |
"every_snd s = SCons (shd s) (stl (stl s))" |
1763 |
||
1764 |
text {*
|
|
| 53752 | 1765 |
\noindent |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1766 |
Constructs such as @{text "let"}---@{text "in"}, @{text
|
| 53646 | 1767 |
"if"}---@{text "then"}---@{text "else"}, and @{text "case"}---@{text "of"} may
|
1768 |
appear around constructors that guard corecursive calls: |
|
1769 |
*} |
|
1770 |
||
| 54072 | 1771 |
primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where |
| 53644 | 1772 |
"lappend xs ys = |
1773 |
(case xs of |
|
1774 |
LNil \<Rightarrow> ys |
|
1775 |
| LCons x xs' \<Rightarrow> LCons x (lappend xs' ys))" |
|
1776 |
||
| 53646 | 1777 |
text {*
|
| 53752 | 1778 |
\noindent |
| 54402 | 1779 |
Pattern matching is not supported by @{command primcorec}. Fortunately, it is
|
1780 |
easy to generate pattern-maching equations using the \keyw{simps\_of\_case}
|
|
1781 |
command provided by the theory \verb|~~/src/HOL/Library/Simps_Case_Conv|. |
|
1782 |
*} |
|
1783 |
||
1784 |
simps_of_case lappend_simps: lappend.code |
|
1785 |
||
1786 |
text {*
|
|
1787 |
This generates the lemma collection @{thm [source] lappend_simps}:
|
|
1788 |
% |
|
1789 |
\[@{thm lappend_simps(1)[no_vars]}
|
|
1790 |
\qquad @{thm lappend_simps(2)[no_vars]}\]
|
|
1791 |
% |
|
| 53646 | 1792 |
Corecursion is useful to specify not only functions but also infinite objects: |
1793 |
*} |
|
1794 |
||
| 53826 | 1795 |
primcorec infty :: enat where |
| 53644 | 1796 |
"infty = ESuc infty" |
1797 |
||
| 53646 | 1798 |
text {*
|
| 53752 | 1799 |
\noindent |
1800 |
The example below constructs a pseudorandom process value. It takes a stream of |
|
| 53675 | 1801 |
actions (@{text s}), a pseudorandom function generator (@{text f}), and a
|
1802 |
pseudorandom seed (@{text n}):
|
|
1803 |
*} |
|
1804 |
||
| 54072 | 1805 |
primcorec |
| 53752 | 1806 |
random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process" |
1807 |
where |
|
| 53675 | 1808 |
"random_process s f n = |
1809 |
(if n mod 4 = 0 then |
|
1810 |
Fail |
|
1811 |
else if n mod 4 = 1 then |
|
1812 |
Skip (random_process s f (f n)) |
|
1813 |
else if n mod 4 = 2 then |
|
1814 |
Action (shd s) (random_process (stl s) f (f n)) |
|
1815 |
else |
|
1816 |
Choice (random_process (every_snd s) (f \<circ> f) (f n)) |
|
1817 |
(random_process (every_snd (stl s)) (f \<circ> f) (f (f n))))" |
|
1818 |
||
1819 |
text {*
|
|
1820 |
\noindent |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1821 |
The main disadvantage of the code view is that the conditions are tested |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1822 |
sequentially. This is visible in the generated theorems. The constructor and |
| 53752 | 1823 |
destructor views offer nonsequential alternatives. |
| 53675 | 1824 |
*} |
1825 |
||
| 53644 | 1826 |
|
1827 |
subsubsection {* Mutual Corecursion
|
|
1828 |
\label{sssec:primcorec-mutual-corecursion} *}
|
|
1829 |
||
| 53647 | 1830 |
text {*
|
1831 |
The syntax for mutually corecursive functions over mutually corecursive |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1832 |
datatypes is unsurprising: |
| 53647 | 1833 |
*} |
1834 |
||
| 53826 | 1835 |
primcorec |
| 53644 | 1836 |
even_infty :: even_enat and |
1837 |
odd_infty :: odd_enat |
|
1838 |
where |
|
1839 |
"even_infty = Even_ESuc odd_infty" | |
|
1840 |
"odd_infty = Odd_ESuc even_infty" |
|
1841 |
||
1842 |
||
1843 |
subsubsection {* Nested Corecursion
|
|
1844 |
\label{sssec:primcorec-nested-corecursion} *}
|
|
1845 |
||
| 53647 | 1846 |
text {*
|
1847 |
The next pair of examples generalize the @{const literate} and @{const siterate}
|
|
1848 |
functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly
|
|
1849 |
infinite trees in which subnodes are organized either as a lazy list (@{text
|
|
| 54072 | 1850 |
tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}). They rely on the map functions of
|
1851 |
the nesting type constructors to lift the corecursive calls: |
|
| 53647 | 1852 |
*} |
1853 |
||
| 53826 | 1854 |
primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
|
| 54072 | 1855 |
"iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i g) (g x))" |
| 53644 | 1856 |
|
| 53677 | 1857 |
text {* \blankline *}
|
1858 |
||
| 53826 | 1859 |
primcorec iterate\<^sub>i\<^sub>s :: "('a \<Rightarrow> 'a fset) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>s" where
|
| 54072 | 1860 |
"iterate\<^sub>i\<^sub>s g x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s g) (g x))" |
| 53644 | 1861 |
|
| 52805 | 1862 |
text {*
|
| 53752 | 1863 |
\noindent |
| 54072 | 1864 |
Both examples follow the usual format for constructor arguments associated |
1865 |
with nested recursive occurrences of the datatype. Consider |
|
1866 |
@{const iterate\<^sub>i\<^sub>i}. The term @{term "g x"} constructs an @{typ "'a llist"}
|
|
1867 |
value, which is turned into an @{typ "'a tree\<^sub>i\<^sub>i llist"} value using
|
|
1868 |
@{const lmap}.
|
|
1869 |
||
1870 |
This format may sometimes feel artificial. The following function constructs |
|
1871 |
a tree with a single, infinite branch from a stream: |
|
1872 |
*} |
|
1873 |
||
1874 |
primcorec tree\<^sub>i\<^sub>i_of_stream :: "'a stream \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where |
|
1875 |
"tree\<^sub>i\<^sub>i_of_stream s = |
|
1876 |
Node\<^sub>i\<^sub>i (shd s) (lmap tree\<^sub>i\<^sub>i_of_stream (LCons (stl s) LNil))" |
|
1877 |
||
1878 |
text {*
|
|
1879 |
\noindent |
|
| 54278 | 1880 |
Fortunately, it is easy to prove the following lemma, where the corecursive call |
1881 |
is moved inside the lazy list constructor, thereby eliminating the need for |
|
1882 |
@{const lmap}:
|
|
| 54072 | 1883 |
*} |
1884 |
||
| 54278 | 1885 |
lemma tree\<^sub>i\<^sub>i_of_stream_alt: |
| 54072 | 1886 |
"tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (LCons (tree\<^sub>i\<^sub>i_of_stream (stl s)) LNil)" |
| 54278 | 1887 |
by (subst tree\<^sub>i\<^sub>i_of_stream.code) simp |
| 54072 | 1888 |
|
1889 |
text {*
|
|
1890 |
The next example illustrates corecursion through functions, which is a bit |
|
1891 |
special. Deterministic finite automata (DFAs) are traditionally defined as |
|
1892 |
5-tuples @{text "(Q, \<Sigma>, \<delta>, q\<^sub>0, F)"}, where @{text Q} is a finite set of states,
|
|
| 53675 | 1893 |
@{text \<Sigma>} is a finite alphabet, @{text \<delta>} is a transition function, @{text q\<^sub>0}
|
1894 |
is an initial state, and @{text F} is a set of final states. The following
|
|
1895 |
function translates a DFA into a @{type state_machine}:
|
|
1896 |
*} |
|
1897 |
||
| 54071 | 1898 |
primcorec |
1899 |
(*<*)(in early) (*>*)sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
|
|
| 53752 | 1900 |
where |
| 54182 | 1901 |
"sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F \<circ> \<delta> q)" |
| 53675 | 1902 |
|
| 53751 | 1903 |
text {*
|
1904 |
\noindent |
|
1905 |
The map function for the function type (@{text \<Rightarrow>}) is composition
|
|
| 54181 | 1906 |
(@{text "op \<circ>"}). For convenience, corecursion through functions can
|
| 54182 | 1907 |
also be expressed using $\lambda$-abstractions and function application rather |
| 54031 | 1908 |
than through composition. For example: |
| 53751 | 1909 |
*} |
1910 |
||
| 53826 | 1911 |
primcorec |
| 53752 | 1912 |
sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
|
1913 |
where |
|
| 54182 | 1914 |
"sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (\<lambda>a. sm_of_dfa \<delta> F (\<delta> q a))" |
| 53752 | 1915 |
|
1916 |
text {* \blankline *}
|
|
1917 |
||
| 53826 | 1918 |
primcorec empty_sm :: "'a state_machine" where |
| 53752 | 1919 |
"empty_sm = State_Machine False (\<lambda>_. empty_sm)" |
| 53751 | 1920 |
|
| 53752 | 1921 |
text {* \blankline *}
|
1922 |
||
| 53826 | 1923 |
primcorec not_sm :: "'a state_machine \<Rightarrow> 'a state_machine" where |
| 53752 | 1924 |
"not_sm M = State_Machine (\<not> accept M) (\<lambda>a. not_sm (trans M a))" |
| 53751 | 1925 |
|
| 53752 | 1926 |
text {* \blankline *}
|
1927 |
||
| 53826 | 1928 |
primcorec |
| 53752 | 1929 |
or_sm :: "'a state_machine \<Rightarrow> 'a state_machine \<Rightarrow> 'a state_machine" |
1930 |
where |
|
| 54072 | 1931 |
"or_sm M N = State_Machine (accept M \<or> accept N) |
1932 |
(\<lambda>a. or_sm (trans M a) (trans N a))" |
|
| 53751 | 1933 |
|
| 54182 | 1934 |
text {*
|
1935 |
\noindent |
|
1936 |
For recursion through curried $n$-ary functions, $n$ applications of |
|
1937 |
@{term "op \<circ>"} are necessary. The examples below illustrate the case where
|
|
1938 |
$n = 2$: |
|
1939 |
*} |
|
1940 |
||
1941 |
codatatype ('a, 'b) state_machine2 =
|
|
1942 |
State_Machine2 (accept2: bool) (trans2: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) state_machine2")
|
|
1943 |
||
1944 |
text {* \blankline *}
|
|
1945 |
||
1946 |
primcorec |
|
1947 |
(*<*)(in early) (*>*)sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) state_machine2"
|
|
1948 |
where |
|
1949 |
"sm2_of_dfa \<delta> F q = State_Machine2 (q \<in> F) (op \<circ> (op \<circ> (sm2_of_dfa \<delta> F)) (\<delta> q))" |
|
1950 |
||
1951 |
text {* \blankline *}
|
|
1952 |
||
1953 |
primcorec |
|
1954 |
sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) state_machine2"
|
|
1955 |
where |
|
1956 |
"sm2_of_dfa \<delta> F q = State_Machine2 (q \<in> F) (\<lambda>a b. sm2_of_dfa \<delta> F (\<delta> q a b))" |
|
1957 |
||
| 53644 | 1958 |
|
1959 |
subsubsection {* Nested-as-Mutual Corecursion
|
|
1960 |
\label{sssec:primcorec-nested-as-mutual-corecursion} *}
|
|
1961 |
||
| 53647 | 1962 |
text {*
|
1963 |
Just as it is possible to recurse over nested recursive datatypes as if they |
|
1964 |
were mutually recursive |
|
1965 |
(Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to
|
|
| 53752 | 1966 |
pretend that nested codatatypes are mutually corecursive. For example: |
| 53647 | 1967 |
*} |
1968 |
||
| 54287 | 1969 |
(*<*) |
1970 |
context late |
|
1971 |
begin |
|
1972 |
(*>*) |
|
| 54072 | 1973 |
primcorec |
| 54287 | 1974 |
iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
|
| 53644 | 1975 |
iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
|
1976 |
where |
|
| 54072 | 1977 |
"iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" | |
1978 |
"iterates\<^sub>i\<^sub>i g xs = |
|
| 53644 | 1979 |
(case xs of |
1980 |
LNil \<Rightarrow> LNil |
|
| 54072 | 1981 |
| LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))" |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1982 |
|
| 54287 | 1983 |
text {*
|
1984 |
\noindent |
|
1985 |
Coinduction rules are generated as |
|
1986 |
@{thm [source] iterate\<^sub>i\<^sub>i.coinduct},
|
|
1987 |
@{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and
|
|
1988 |
@{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct}
|
|
1989 |
and analogously for @{text strong_coinduct}. These rules and the
|
|
1990 |
underlying corecursors are generated on a per-need basis and are kept in a cache |
|
1991 |
to speed up subsequent definitions. |
|
1992 |
*} |
|
1993 |
||
1994 |
(*<*) |
|
1995 |
end |
|
1996 |
(*>*) |
|
1997 |
||
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1998 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
1999 |
subsubsection {* Constructor View
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2000 |
\label{ssec:primrec-constructor-view} *}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2001 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2002 |
(*<*) |
| 54182 | 2003 |
locale ctr_view |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2004 |
begin |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2005 |
(*>*) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2006 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2007 |
text {*
|
| 53750 | 2008 |
The constructor view is similar to the code view, but there is one separate |
2009 |
conditional equation per constructor rather than a single unconditional |
|
2010 |
equation. Examples that rely on a single constructor, such as @{const literate}
|
|
2011 |
and @{const siterate}, are identical in both styles.
|
|
2012 |
||
2013 |
Here is an example where there is a difference: |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2014 |
*} |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2015 |
|
| 53826 | 2016 |
primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2017 |
"lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lappend xs ys = LNil" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2018 |
"_ \<Longrightarrow> lappend xs ys = LCons (lhd (if lnull xs then ys else xs)) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2019 |
(if xs = LNil then ltl ys else lappend (ltl xs) ys)" |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2020 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2021 |
text {*
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2022 |
\noindent |
| 53752 | 2023 |
With the constructor view, we must distinguish between the @{const LNil} and
|
2024 |
the @{const LCons} case. The condition for @{const LCons} is
|
|
2025 |
left implicit, as the negation of that for @{const LNil}.
|
|
| 53750 | 2026 |
|
2027 |
For this example, the constructor view is slighlty more involved than the |
|
2028 |
code equation. Recall the code view version presented in |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2029 |
Section~\ref{sssec:primcorec-simple-corecursion}.
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2030 |
% TODO: \[{thm code_view.lappend.code}\]
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2031 |
The constructor view requires us to analyze the second argument (@{term ys}).
|
| 53752 | 2032 |
The code equation generated from the constructor view also suffers from this. |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2033 |
% TODO: \[{thm lappend.code}\]
|
| 53750 | 2034 |
|
| 53752 | 2035 |
In contrast, the next example is arguably more naturally expressed in the |
2036 |
constructor view: |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2037 |
*} |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2038 |
|
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53829
diff
changeset
|
2039 |
primcorec |
| 53752 | 2040 |
random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process" |
2041 |
where |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2042 |
"n mod 4 = 0 \<Longrightarrow> random_process s f n = Fail" | |
| 53752 | 2043 |
"n mod 4 = 1 \<Longrightarrow> |
2044 |
random_process s f n = Skip (random_process s f (f n))" | |
|
2045 |
"n mod 4 = 2 \<Longrightarrow> |
|
2046 |
random_process s f n = Action (shd s) (random_process (stl s) f (f n))" | |
|
2047 |
"n mod 4 = 3 \<Longrightarrow> |
|
2048 |
random_process s f n = Choice (random_process (every_snd s) f (f n)) |
|
| 53826 | 2049 |
(random_process (every_snd (stl s)) f (f n))" |
2050 |
(*<*) |
|
| 53644 | 2051 |
end |
2052 |
(*>*) |
|
| 52805 | 2053 |
|
| 53750 | 2054 |
text {*
|
| 53752 | 2055 |
\noindent |
| 53750 | 2056 |
Since there is no sequentiality, we can apply the equation for @{const Choice}
|
| 53752 | 2057 |
without having first to discharge @{term "n mod (4\<Colon>int) \<noteq> 0"},
|
2058 |
@{term "n mod (4\<Colon>int) \<noteq> 1"}, and
|
|
2059 |
@{term "n mod (4\<Colon>int) \<noteq> 2"}.
|
|
| 53750 | 2060 |
The price to pay for this elegance is that we must discharge exclusivity proof |
2061 |
obligations, one for each pair of conditions |
|
| 53752 | 2062 |
@{term "(n mod (4\<Colon>int) = i, n mod (4\<Colon>int) = j)"}
|
2063 |
with @{term "i < j"}. If we prefer not to discharge any obligations, we can
|
|
2064 |
enable the @{text "sequential"} option. This pushes the problem to the users of
|
|
2065 |
the generated properties. |
|
| 53750 | 2066 |
%Here are more examples to conclude: |
2067 |
*} |
|
2068 |
||
| 52824 | 2069 |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2070 |
subsubsection {* Destructor View
|
| 53752 | 2071 |
\label{ssec:primrec-destructor-view} *}
|
2072 |
||
2073 |
(*<*) |
|
| 54182 | 2074 |
locale dtr_view |
| 53752 | 2075 |
begin |
2076 |
(*>*) |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2077 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2078 |
text {*
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2079 |
The destructor view is in many respects dual to the constructor view. Conditions |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2080 |
determine which constructor to choose, and these conditions are interpreted |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2081 |
sequentially or not depending on the @{text "sequential"} option.
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2082 |
Consider the following examples: |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2083 |
*} |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2084 |
|
| 53826 | 2085 |
primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2086 |
"\<not> lnull (literate _ x)" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2087 |
"lhd (literate _ x) = x" | |
| 54072 | 2088 |
"ltl (literate g x) = literate g (g x)" |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2089 |
|
| 53752 | 2090 |
text {* \blankline *}
|
2091 |
||
| 53826 | 2092 |
primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2093 |
"shd (siterate _ x) = x" | |
| 54072 | 2094 |
"stl (siterate g x) = siterate g (g x)" |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2095 |
|
| 53752 | 2096 |
text {* \blankline *}
|
2097 |
||
| 53826 | 2098 |
primcorec every_snd :: "'a stream \<Rightarrow> 'a stream" where |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2099 |
"shd (every_snd s) = shd s" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2100 |
"stl (every_snd s) = stl (stl s)" |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2101 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2102 |
text {*
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2103 |
\noindent |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2104 |
The first formula in the @{const literate} specification indicates which
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2105 |
constructor to choose. For @{const siterate} and @{const every_snd}, no such
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2106 |
formula is necessary, since the type has only one constructor. The last two |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2107 |
formulas are equations specifying the value of the result for the relevant |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2108 |
selectors. Corecursive calls appear directly to the right of the equal sign. |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2109 |
Their arguments are unrestricted. |
| 53750 | 2110 |
|
2111 |
The next example shows how to specify functions that rely on more than one |
|
2112 |
constructor: |
|
2113 |
*} |
|
2114 |
||
| 53826 | 2115 |
primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where |
| 53750 | 2116 |
"lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" | |
2117 |
"lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" | |
|
2118 |
"ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)" |
|
2119 |
||
2120 |
text {*
|
|
2121 |
\noindent |
|
2122 |
For a codatatype with $n$ constructors, it is sufficient to specify $n - 1$ |
|
2123 |
discriminator formulas. The command will then assume that the remaining |
|
2124 |
constructor should be taken otherwise. This can be made explicit by adding |
|
2125 |
*} |
|
2126 |
||
2127 |
(*<*) |
|
2128 |
end |
|
2129 |
||
| 54182 | 2130 |
locale dtr_view2 |
2131 |
begin |
|
2132 |
||
| 53826 | 2133 |
primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where |
| 53750 | 2134 |
"lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" | |
2135 |
(*>*) |
|
| 53752 | 2136 |
"_ \<Longrightarrow> \<not> lnull (lappend xs ys)" |
2137 |
(*<*) | |
|
| 53750 | 2138 |
"lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" | |
2139 |
"ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)" |
|
2140 |
(*>*) |
|
2141 |
||
2142 |
text {*
|
|
2143 |
\noindent |
|
| 53752 | 2144 |
to the specification. The generated selector theorems are conditional. |
2145 |
||
2146 |
The next example illustrates how to cope with selectors defined for several |
|
| 53750 | 2147 |
constructors: |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2148 |
*} |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2149 |
|
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53829
diff
changeset
|
2150 |
primcorec |
| 53752 | 2151 |
random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process" |
2152 |
where |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2153 |
"n mod 4 = 0 \<Longrightarrow> is_Fail (random_process s f n)" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2154 |
"n mod 4 = 1 \<Longrightarrow> is_Skip (random_process s f n)" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2155 |
"n mod 4 = 2 \<Longrightarrow> is_Action (random_process s f n)" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2156 |
"n mod 4 = 3 \<Longrightarrow> is_Choice (random_process s f n)" | |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53829
diff
changeset
|
2157 |
"cont (random_process s f n) = random_process s f (f n)" of Skip | |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2158 |
"prefix (random_process s f n) = shd s" | |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53829
diff
changeset
|
2159 |
"cont (random_process s f n) = random_process (stl s) f (f n)" of Action | |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2160 |
"left (random_process s f n) = random_process (every_snd s) f (f n)" | |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53829
diff
changeset
|
2161 |
"right (random_process s f n) = random_process (every_snd (stl s)) f (f n)" |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2162 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2163 |
text {*
|
| 53750 | 2164 |
\noindent |
2165 |
Using the @{text "of"} keyword, different equations are specified for @{const
|
|
2166 |
cont} depending on which constructor is selected. |
|
2167 |
||
2168 |
Here are more examples to conclude: |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2169 |
*} |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2170 |
|
| 53826 | 2171 |
primcorec |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2172 |
even_infty :: even_enat and |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2173 |
odd_infty :: odd_enat |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2174 |
where |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2175 |
"\<not> is_Even_EZero even_infty" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2176 |
"un_Even_ESuc even_infty = odd_infty" | |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2177 |
"un_Odd_ESuc odd_infty = even_infty" |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2178 |
|
| 53752 | 2179 |
text {* \blankline *}
|
2180 |
||
| 53826 | 2181 |
primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
|
| 54072 | 2182 |
"lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = x" | |
2183 |
"sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = lmap (iterate\<^sub>i\<^sub>i g) (g x)" |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2184 |
(*<*) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2185 |
end |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2186 |
(*>*) |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2187 |
|
| 53750 | 2188 |
|
| 53617 | 2189 |
subsection {* Command Syntax
|
2190 |
\label{ssec:primcorec-command-syntax} *}
|
|
2191 |
||
2192 |
||
| 53826 | 2193 |
subsubsection {* \keyw{primcorec} and \keyw{primcorecursive}
|
|
53753
ae7f50e70c09
renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents:
53752
diff
changeset
|
2194 |
\label{sssec:primcorecursive-and-primcorec} *}
|
| 52840 | 2195 |
|
2196 |
text {*
|
|
| 53829 | 2197 |
\begin{matharray}{rcl}
|
2198 |
@{command_def "primcorec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
|
|
2199 |
@{command_def "primcorecursive"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
|
|
2200 |
\end{matharray}
|
|
| 52840 | 2201 |
|
2202 |
@{rail "
|
|
| 53829 | 2203 |
(@@{command primcorec} | @@{command primcorecursive}) target? \\ @{syntax pcr_option}? fixes @'where'
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2204 |
(@{syntax pcr_formula} + '|')
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2205 |
; |
| 53828 | 2206 |
@{syntax_def pcr_option}: '(' ('sequential' | 'exhaustive') ')'
|
| 52840 | 2207 |
; |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2208 |
@{syntax_def pcr_formula}: thmdecl? prop (@'of' (term * ))?
|
| 52840 | 2209 |
"} |
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2210 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2211 |
The optional target is optionally followed by a corecursion-specific option: |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2212 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2213 |
\begin{itemize}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2214 |
\setlength{\itemsep}{0pt}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2215 |
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2216 |
\item |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2217 |
The @{text "sequential"} option indicates that the conditions in specifications
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2218 |
expressed using the constructor or destructor view are to be interpreted |
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2219 |
sequentially. |
| 53826 | 2220 |
|
2221 |
\item |
|
2222 |
The @{text "exhaustive"} option indicates that the conditions in specifications
|
|
2223 |
expressed using the constructor or destructor view cover all possible cases. |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2224 |
\end{itemize}
|
|
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2225 |
|
| 53826 | 2226 |
\noindent |
2227 |
The @{command primcorec} command is an abbreviation for @{command primcorecursive} with
|
|
2228 |
@{text "by auto?"} to discharge any emerging proof obligations.
|
|
| 52840 | 2229 |
*} |
| 52794 | 2230 |
|
| 52824 | 2231 |
|
| 53619 | 2232 |
(* |
| 52840 | 2233 |
subsection {* Generated Theorems
|
2234 |
\label{ssec:primcorec-generated-theorems} *}
|
|
| 53619 | 2235 |
*) |
| 52794 | 2236 |
|
2237 |
||
| 53623 | 2238 |
(* |
2239 |
subsection {* Recursive Default Values for Selectors
|
|
2240 |
\label{ssec:primcorec-recursive-default-values-for-selectors} *}
|
|
2241 |
||
2242 |
text {*
|
|
2243 |
partial_function to the rescue |
|
2244 |
*} |
|
2245 |
*) |
|
2246 |
||
2247 |
||
| 52827 | 2248 |
section {* Registering Bounded Natural Functors
|
| 52805 | 2249 |
\label{sec:registering-bounded-natural-functors} *}
|
| 52792 | 2250 |
|
| 52805 | 2251 |
text {*
|
| 53647 | 2252 |
The (co)datatype package can be set up to allow nested recursion through |
2253 |
arbitrary type constructors, as long as they adhere to the BNF requirements and |
|
2254 |
are registered as BNFs. |
|
| 52805 | 2255 |
*} |
2256 |
||
| 52824 | 2257 |
|
| 53619 | 2258 |
(* |
| 53617 | 2259 |
subsection {* Introductory Example
|
2260 |
\label{ssec:bnf-introductory-example} *}
|
|
| 52805 | 2261 |
|
2262 |
text {*
|
|
2263 |
More examples in \verb|~~/src/HOL/BNF/Basic_BNFs.thy| and |
|
2264 |
\verb|~~/src/HOL/BNF/More_BNFs.thy|. |
|
| 52806 | 2265 |
|
| 53617 | 2266 |
%Mention distinction between live and dead type arguments; |
2267 |
% * and existence of map, set for those |
|
2268 |
%mention =>. |
|
| 52805 | 2269 |
*} |
| 53619 | 2270 |
*) |
| 52794 | 2271 |
|
| 52824 | 2272 |
|
| 53617 | 2273 |
subsection {* Command Syntax
|
2274 |
\label{ssec:bnf-command-syntax} *}
|
|
2275 |
||
2276 |
||
| 53621 | 2277 |
subsubsection {* \keyw{bnf}
|
2278 |
\label{sssec:bnf} *}
|
|
| 52794 | 2279 |
|
| 53028 | 2280 |
text {*
|
| 53829 | 2281 |
\begin{matharray}{rcl}
|
2282 |
@{command_def "bnf"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
|
|
2283 |
\end{matharray}
|
|
2284 |
||
| 53028 | 2285 |
@{rail "
|
| 54421 | 2286 |
@@{command bnf} target? (name ':')? typ \\
|
2287 |
'map:' term ('sets:' (term +))? 'bd:' term \\
|
|
2288 |
('wits:' (term +))? ('rel:' term)?
|
|
| 53028 | 2289 |
"} |
2290 |
*} |
|
| 52805 | 2291 |
|
| 53617 | 2292 |
|
| 54211 | 2293 |
(* NOTYET |
| 54187 | 2294 |
subsubsection {* \keyw{bnf\_decl}
|
2295 |
\label{sssec:bnf-decl} *}
|
|
2296 |
||
2297 |
text {*
|
|
2298 |
%%% TODO: use command_def once the command is available |
|
2299 |
\begin{matharray}{rcl}
|
|
2300 |
@{text "bnf_decl"} & : & @{text "local_theory \<rightarrow> local_theory"}
|
|
2301 |
\end{matharray}
|
|
2302 |
||
2303 |
@{rail "
|
|
2304 |
@@{command bnf} target? dt_name
|
|
2305 |
"} |
|
2306 |
*} |
|
| 54211 | 2307 |
*) |
| 54187 | 2308 |
|
2309 |
||
| 53621 | 2310 |
subsubsection {* \keyw{print\_bnfs}
|
2311 |
\label{sssec:print-bnfs} *}
|
|
| 53617 | 2312 |
|
2313 |
text {*
|
|
| 53829 | 2314 |
\begin{matharray}{rcl}
|
2315 |
@{command_def "print_bnfs"} & : & @{text "local_theory \<rightarrow>"}
|
|
2316 |
\end{matharray}
|
|
2317 |
||
| 53647 | 2318 |
@{rail "
|
| 53829 | 2319 |
@@{command print_bnfs}
|
| 53647 | 2320 |
"} |
| 53617 | 2321 |
*} |
2322 |
||
2323 |
||
2324 |
section {* Deriving Destructors and Theorems for Free Constructors
|
|
2325 |
\label{sec:deriving-destructors-and-theorems-for-free-constructors} *}
|
|
| 52794 | 2326 |
|
| 52805 | 2327 |
text {*
|
| 53623 | 2328 |
The derivation of convenience theorems for types equipped with free constructors, |
| 53829 | 2329 |
as performed internally by @{command datatype_new} and @{command codatatype},
|
| 53623 | 2330 |
is available as a stand-alone command called @{command wrap_free_constructors}.
|
| 52794 | 2331 |
|
| 53617 | 2332 |
% * need for this is rare but may arise if you want e.g. to add destructors to |
2333 |
% a type not introduced by ... |
|
2334 |
% |
|
2335 |
% * also useful for compatibility with old package, e.g. add destructors to |
|
2336 |
% old \keyw{datatype}
|
|
2337 |
% |
|
2338 |
% * @{command wrap_free_constructors}
|
|
| 53623 | 2339 |
% * @{text "no_discs_sels"}, @{text "rep_compat"}
|
| 53617 | 2340 |
% * hack to have both co and nonco view via locale (cf. ext nats) |
| 52805 | 2341 |
*} |
| 52792 | 2342 |
|
| 52824 | 2343 |
|
| 53619 | 2344 |
(* |
| 53617 | 2345 |
subsection {* Introductory Example
|
2346 |
\label{ssec:ctors-introductory-example} *}
|
|
| 53619 | 2347 |
*) |
| 52794 | 2348 |
|
| 52824 | 2349 |
|
| 53617 | 2350 |
subsection {* Command Syntax
|
2351 |
\label{ssec:ctors-command-syntax} *}
|
|
2352 |
||
2353 |
||
| 53621 | 2354 |
subsubsection {* \keyw{wrap\_free\_constructors}
|
| 53675 | 2355 |
\label{sssec:wrap-free-constructors} *}
|
| 52828 | 2356 |
|
| 53018 | 2357 |
text {*
|
| 53829 | 2358 |
\begin{matharray}{rcl}
|
2359 |
@{command_def "wrap_free_constructors"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
|
|
2360 |
\end{matharray}
|
|
| 53018 | 2361 |
|
2362 |
@{rail "
|
|
| 53829 | 2363 |
@@{command wrap_free_constructors} target? @{syntax dt_options} \\
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
2364 |
term_list name @{syntax wfc_discs_sels}?
|
| 53018 | 2365 |
; |
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
2366 |
@{syntax_def wfc_discs_sels}: name_list (name_list_list name_term_list_list? )?
|
| 53018 | 2367 |
; |
| 53534 | 2368 |
@{syntax_def name_term}: (name ':' term)
|
| 54421 | 2369 |
; |
2370 |
X_list: '[' (X + ',') ']' |
|
| 53018 | 2371 |
"} |
2372 |
||
| 53617 | 2373 |
% options: no_discs_sels rep_compat |
| 53028 | 2374 |
|
| 53617 | 2375 |
% X_list is as for BNF |
| 53028 | 2376 |
|
| 53829 | 2377 |
\noindent |
| 53542 | 2378 |
Section~\ref{ssec:datatype-generated-theorems} lists the generated theorems.
|
| 53018 | 2379 |
*} |
| 52828 | 2380 |
|
| 52794 | 2381 |
|
| 53617 | 2382 |
(* |
| 52827 | 2383 |
section {* Standard ML Interface
|
| 52805 | 2384 |
\label{sec:standard-ml-interface} *}
|
| 52792 | 2385 |
|
| 52805 | 2386 |
text {*
|
| 53623 | 2387 |
The package's programmatic interface. |
| 52805 | 2388 |
*} |
| 53617 | 2389 |
*) |
| 52794 | 2390 |
|
2391 |
||
| 53617 | 2392 |
(* |
| 52827 | 2393 |
section {* Interoperability
|
| 52805 | 2394 |
\label{sec:interoperability} *}
|
| 52794 | 2395 |
|
| 52805 | 2396 |
text {*
|
| 53623 | 2397 |
The package's interaction with other Isabelle packages and tools, such as the |
2398 |
code generator and the counterexample generators. |
|
| 52805 | 2399 |
*} |
| 52794 | 2400 |
|
| 52824 | 2401 |
|
| 52828 | 2402 |
subsection {* Transfer and Lifting
|
2403 |
\label{ssec:transfer-and-lifting} *}
|
|
| 52794 | 2404 |
|
| 52824 | 2405 |
|
| 52828 | 2406 |
subsection {* Code Generator
|
2407 |
\label{ssec:code-generator} *}
|
|
| 52794 | 2408 |
|
| 52824 | 2409 |
|
| 52828 | 2410 |
subsection {* Quickcheck
|
2411 |
\label{ssec:quickcheck} *}
|
|
| 52794 | 2412 |
|
| 52824 | 2413 |
|
| 52828 | 2414 |
subsection {* Nitpick
|
2415 |
\label{ssec:nitpick} *}
|
|
| 52794 | 2416 |
|
| 52824 | 2417 |
|
| 52828 | 2418 |
subsection {* Nominal Isabelle
|
2419 |
\label{ssec:nominal-isabelle} *}
|
|
| 53617 | 2420 |
*) |
| 52794 | 2421 |
|
| 52805 | 2422 |
|
| 53617 | 2423 |
(* |
| 52827 | 2424 |
section {* Known Bugs and Limitations
|
| 52805 | 2425 |
\label{sec:known-bugs-and-limitations} *}
|
2426 |
||
2427 |
text {*
|
|
| 53623 | 2428 |
Known open issues of the package. |
| 52805 | 2429 |
*} |
| 52794 | 2430 |
|
2431 |
text {*
|
|
|
53753
ae7f50e70c09
renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents:
53752
diff
changeset
|
2432 |
%* primcorecursive and primcorec is unfinished |
| 53617 | 2433 |
% |
2434 |
%* slow n-ary mutual (co)datatype, avoid as much as possible (e.g. using nesting) |
|
2435 |
% |
|
2436 |
%* issues with HOL-Proofs? |
|
2437 |
% |
|
2438 |
%* partial documentation |
|
2439 |
% |
|
2440 |
%* no way to register "sum" and "prod" as (co)datatypes to enable N2M reduction for them |
|
2441 |
% (for @{command datatype_new_compat} and prim(co)rec)
|
|
2442 |
% |
|
| 53619 | 2443 |
% * a fortiori, no way to register same type as both data- and codatatype |
| 53617 | 2444 |
% |
2445 |
%* no recursion through unused arguments (unlike with the old package) |
|
2446 |
% |
|
2447 |
%* in a locale, cannot use locally fixed types (because of limitation in typedef)? |
|
| 53619 | 2448 |
% |
2449 |
% *names of variables suboptimal |
|
| 52822 | 2450 |
*} |
| 53675 | 2451 |
*) |
| 52822 | 2452 |
|
2453 |
||
2454 |
text {*
|
|
|
53863
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
2455 |
\section*{Acknowledgment}
|
|
c7364dca96f2
textual improvements following Christian Sternagel's feedback
blanchet
parents:
53857
diff
changeset
|
2456 |
|
|
53749
b37db925b663
adapted primcorec documentation to reflect the three views
blanchet
parents:
53748
diff
changeset
|
2457 |
Tobias Nipkow and Makarius Wenzel encouraged us to implement the new |
| 53617 | 2458 |
(co)datatype package. Andreas Lochbihler provided lots of comments on earlier |
2459 |
versions of the package, especially for the coinductive part. Brian Huffman |
|
2460 |
suggested major simplifications to the internal constructions, much of which has |
|
2461 |
yet to be implemented. Florian Haftmann and Christian Urban provided general |
|
| 53675 | 2462 |
advice on Isabelle and package writing. Stefan Milius and Lutz Schr\"oder |
| 54146 | 2463 |
found an elegant proof to eliminate one of the BNF assumptions. Andreas |
2464 |
Lochbihler and Christian Sternagel suggested many textual improvements to this |
|
2465 |
tutorial. |
|
| 52794 | 2466 |
*} |
| 53617 | 2467 |
|
| 52792 | 2468 |
end |