src/Doc/Datatypes/Datatypes.thy
author blanchet
Sun, 15 Sep 2013 23:02:23 +0200
changeset 53642 05ca82603671
parent 53623 501e2091182b
child 53643 673cb2c9d695
permissions -rw-r--r--
more (co)data docs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     1
(*  Title:      Doc/Datatypes/Datatypes.thy
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
     3
    Author:     Lorenz Panny, TU Muenchen
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
     4
    Author:     Andrei Popescu, TU Muenchen
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
     5
    Author:     Dmitriy Traytel, TU Muenchen
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     6
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     7
Tutorial for (co)datatype definitions with the new package.
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     8
*)
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     9
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    10
theory Datatypes
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
    11
imports Setup
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    12
begin
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    13
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    14
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    15
section {* Introduction
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    16
  \label{sec:introduction} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    17
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    18
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    19
The 2013 edition of Isabelle introduced a new definitional package for freely
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    20
generated datatypes and codatatypes. The datatype support is similar to that
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    21
provided by the earlier package due to Berghofer and Wenzel
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    22
\cite{Berghofer-Wenzel:1999:TPHOL}, documented in the Isar reference manual
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
    23
\cite{isabelle-isar-ref}; indeed, replacing the keyword \keyw{datatype} by
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    24
@{command datatype_new} is usually all that is needed to port existing theories
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    25
to use the new package.
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    26
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    27
Perhaps the main advantage of the new package is that it supports recursion
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
    28
through a large class of non-datatypes, such as finite sets:
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    29
*}
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    30
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    31
    datatype_new 'a tree\<^sub>f\<^sub>s = Node\<^sub>f\<^sub>s 'a "'a tree\<^sub>f\<^sub>s fset"
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    32
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    33
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    34
\noindent
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
    35
Another strong point is the support for local definitions:
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    36
*}
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    37
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    38
    context linorder
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    39
    begin
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    40
    datatype_new flag = Less | Eq | Greater
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    41
    end
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    42
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    43
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    44
\noindent
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    45
The package also provides some convenience, notably automatically generated
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
    46
discriminators and selectors.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    47
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
    48
In addition to plain inductive datatypes, the new package supports coinductive
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    49
datatypes, or \emph{codatatypes}, which may have infinite values. For example,
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    50
the following command introduces the type of lazy lists, which comprises both
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    51
finite and infinite values:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    52
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    53
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    54
(*<*)
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    55
    locale dummy_llist
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    56
    begin
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    57
(*>*)
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    58
    codatatype 'a llist = LNil | LCons 'a "'a llist"
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    59
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    60
text {*
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    61
\noindent
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    62
Mixed inductive--coinductive recursion is possible via nesting. Compare the
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    63
following four Rose tree examples:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    64
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    65
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    66
    datatype_new 'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f 'a "'a tree\<^sub>f\<^sub>f list"
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    67
    datatype_new 'a tree\<^sub>f\<^sub>i = Node\<^sub>f\<^sub>i 'a "'a tree\<^sub>f\<^sub>i llist"
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    68
    codatatype 'a tree\<^sub>i\<^sub>f = Node\<^sub>i\<^sub>f 'a "'a tree\<^sub>i\<^sub>f list"
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    69
    codatatype 'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist"
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
    70
(*<*)
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
    71
    end
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
    72
(*>*)
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    73
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    74
text {*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    75
The first two tree types allow only finite branches, whereas the last two allow
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    76
branches of infinite length. Orthogonally, the nodes in the first and third
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    77
types have finite branching, whereas those of the second and fourth may have
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    78
infinitely many direct subtrees.
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    79
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    80
To use the package, it is necessary to import the @{theory BNF} theory, which
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
    81
can be precompiled into the \texttt{HOL-BNF} image. The following commands show
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
    82
how to launch jEdit/PIDE with the image loaded and how to build the image
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
    83
without launching jEdit:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    84
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    85
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    86
text {*
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    87
\noindent
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
    88
\ \ \ \ \texttt{isabelle jedit -l HOL-BNF} \\
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
    89
\noindent
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    90
\hbox{}\ \ \ \ \texttt{isabelle build -b HOL-BNF}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    91
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    92
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    93
text {*
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    94
The package, like its predecessor, fully adheres to the LCF philosophy
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    95
\cite{mgordon79}: The characteristic theorems associated with the specified
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    96
(co)datatypes are derived rather than introduced axiomatically.%
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
    97
\footnote{If the @{text quick_and_dirty} option is enabled, some of the
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    98
internal constructions and most of the internal proof obligations are skipped.}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    99
The package's metatheory is described in a pair of papers
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   100
\cite{traytel-et-al-2012,blanchette-et-al-wit}. The central notion is that of a
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   101
\emph{bounded natural functor} (BNF)---a well-behaved type constructor for which
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   102
nested (co)recursion is supported.
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   103
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   104
This tutorial is organized as follows:
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   105
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   106
\begin{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   107
\setlength{\itemsep}{0pt}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   108
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   109
\item Section \ref{sec:defining-datatypes}, ``Defining Datatypes,''
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   110
describes how to specify datatypes using the @{command datatype_new} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   111
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   112
\item Section \ref{sec:defining-recursive-functions}, ``Defining Recursive
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   113
Functions,'' describes how to specify recursive functions using
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   114
@{command primrec_new}, \keyw{fun}, and \keyw{function}.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   115
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   116
\item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,''
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   117
describes how to specify codatatypes using the @{command codatatype} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   118
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   119
\item Section \ref{sec:defining-corecursive-functions}, ``Defining Corecursive
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   120
Functions,'' describes how to specify corecursive functions using the
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   121
@{command primcorec} command.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   122
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   123
\item Section \ref{sec:registering-bounded-natural-functors}, ``Registering
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   124
Bounded Natural Functors,'' explains how to use the @{command bnf} command
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   125
to register arbitrary type constructors as BNFs.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   126
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   127
\item Section
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   128
\ref{sec:deriving-destructors-and-theorems-for-free-constructors},
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   129
``Deriving Destructors and Theorems for Free Constructors,'' explains how to
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   130
use the command @{command wrap_free_constructors} to derive destructor constants
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   131
and theorems for freely generated types, as performed internally by @{command
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   132
datatype_new} and @{command codatatype}.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   133
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   134
%\item Section \ref{sec:standard-ml-interface}, ``Standard ML Interface,''
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   135
%describes the package's programmatic interface.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   136
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   137
%\item Section \ref{sec:interoperability}, ``Interoperability,''
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   138
%is concerned with the packages' interaction with other Isabelle packages and
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   139
%tools, such as the code generator and the counterexample generators.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   140
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   141
%\item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   142
%Limitations,'' concludes with known open issues at the time of writing.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   143
\end{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   144
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   145
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   146
\newbox\boxA
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   147
\setbox\boxA=\hbox{\texttt{nospam}}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   148
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   149
\newcommand\authoremaili{\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@\allowbreak
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   150
in.\allowbreak tum.\allowbreak de}}
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   151
\newcommand\authoremailii{\texttt{lore{\color{white}nospam}\kern-\wd\boxA{}nz.panny@\allowbreak
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   152
\allowbreak tum.\allowbreak de}}
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   153
\newcommand\authoremailiii{\texttt{pope{\color{white}nospam}\kern-\wd\boxA{}scua@\allowbreak
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   154
in.\allowbreak tum.\allowbreak de}}
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   155
\newcommand\authoremailiv{\texttt{tray{\color{white}nospam}\kern-\wd\boxA{}tel@\allowbreak
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   156
in.\allowbreak tum.\allowbreak de}}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   157
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   158
The commands @{command datatype_new} and @{command primrec_new} are expected to
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   159
displace \keyw{datatype} and \keyw{primrec} in a future release. Authors of new
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   160
theories are encouraged to use the new commands, and maintainers of older
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   161
theories may want to consider upgrading.
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   162
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   163
Comments and bug reports concerning either the tool or this tutorial should be
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   164
directed to the authors at \authoremaili, \authoremailii, \authoremailiii,
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   165
and \authoremailiv.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   166
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   167
\begin{framed}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   168
\noindent
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   169
\textbf{Warning:} This tutorial is under heavy construction. Please apologise
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   170
for its appearance. If you have ideas regarding material that should be
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   171
included, please let the authors know.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   172
\end{framed}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   173
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   174
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   175
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   176
section {* Defining Datatypes
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   177
  \label{sec:defining-datatypes} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   178
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   179
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   180
Datatypes can be specified using the @{command datatype_new} command. The
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   181
command is first illustrated through concrete examples featuring different
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   182
flavors of recursion. More examples can be found in the directory
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   183
\verb|~~/src/HOL/BNF/Examples|.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   184
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   185
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   186
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   187
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   188
  \label{ssec:datatype-introductory-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   189
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   190
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   191
subsubsection {* Nonrecursive Types
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   192
  \label{sssec:datatype-nonrecursive-types} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   193
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   194
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   195
Datatypes are introduced by specifying the desired names and argument types for
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   196
their constructors. \emph{Enumeration} types are the simplest form of datatype.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   197
All their constructors are nullary:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   198
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   199
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   200
    datatype_new trool = Truue | Faalse | Perhaaps
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   201
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   202
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   203
\noindent
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   204
Here, @{const Truue}, @{const Faalse}, and @{const Perhaaps} have the type @{typ trool}.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   205
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   206
Polymorphic types are possible, such as the following option type, modeled after
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   207
its homologue from the @{theory Option} theory:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   208
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   209
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   210
(*<*)
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   211
    hide_const None Some
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   212
(*>*)
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   213
    datatype_new 'a option = None | Some 'a
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   214
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   215
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   216
\noindent
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   217
The constructors are @{text "None :: 'a option"} and
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   218
@{text "Some :: 'a \<Rightarrow> 'a option"}.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   219
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   220
The next example has three type parameters:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   221
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   222
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   223
    datatype_new ('a, 'b, 'c) triple = Triple 'a 'b 'c
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   224
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   225
text {*
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   226
\noindent
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   227
The constructor is
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   228
@{text "Triple :: 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> ('a, 'b, 'c) triple"}.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   229
Unlike in Standard ML, curried constructors are supported. The uncurried variant
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   230
is also possible:
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   231
*}
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   232
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   233
    datatype_new ('a, 'b, 'c) triple\<^sub>u = Triple\<^sub>u "'a * 'b * 'c"
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   234
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   235
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   236
subsubsection {* Simple Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   237
  \label{sssec:datatype-simple-recursion} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   238
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   239
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   240
Natural numbers are the simplest example of a recursive type:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   241
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   242
53332
c97a05a26dd6 Doc improvements
traytel
parents: 53331
diff changeset
   243
    datatype_new nat = Zero | Suc nat
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   244
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   245
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   246
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   247
Lists were shown in the introduction. Terminated lists are a variant:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   248
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   249
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   250
(*<*)
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   251
    locale dummy_tlist
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   252
    begin
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   253
(*>*)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   254
    datatype_new ('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist"
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   255
(*<*)
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   256
    end
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   257
(*>*)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   258
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   259
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   260
\noindent
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   261
Occurrences of nonatomic types on the right-hand side of the equal sign must be
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   262
enclosed in double quotes, as is customary in Isabelle.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   263
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   264
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   265
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   266
subsubsection {* Mutual Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   267
  \label{sssec:datatype-mutual-recursion} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   268
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   269
text {*
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   270
\emph{Mutually recursive} types are introduced simultaneously and may refer to
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   271
each other. The example below introduces a pair of types for even and odd
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   272
natural numbers:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   273
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   274
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   275
    datatype_new even_nat = Even_Zero | Even_Suc odd_nat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   276
    and odd_nat = Odd_Suc even_nat
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   277
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   278
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   279
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   280
Arithmetic expressions are defined via terms, terms via factors, and factors via
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   281
expressions:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   282
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   283
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   284
    datatype_new ('a, 'b) exp =
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   285
      Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   286
    and ('a, 'b) trm =
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   287
      Factor "('a, 'b) fct" | Prod "('a, 'b) fct" "('a, 'b) trm"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   288
    and ('a, 'b) fct =
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   289
      Const 'a | Var 'b | Expr "('a, 'b) exp"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   290
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   291
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   292
subsubsection {* Nested Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   293
  \label{sssec:datatype-nested-recursion} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   294
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   295
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   296
\emph{Nested recursion} occurs when recursive occurrences of a type appear under
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   297
a type constructor. The introduction showed some examples of trees with nesting
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   298
through lists. A more complex example, that reuses our @{text option} type,
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   299
follows:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   300
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   301
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   302
    datatype_new 'a btree =
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   303
      BNode 'a "'a btree option" "'a btree option"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   304
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   305
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   306
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   307
Not all nestings are admissible. For example, this command will fail:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   308
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   309
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   310
    datatype_new 'a wrong = Wrong (*<*)'a
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   311
    typ (*>*)"'a wrong \<Rightarrow> 'a"
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   312
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   313
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   314
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   315
The issue is that the function arrow @{text "\<Rightarrow>"} allows recursion
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   316
only through its right-hand side. This issue is inherited by polymorphic
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   317
datatypes defined in terms of~@{text "\<Rightarrow>"}:
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   318
*}
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   319
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   320
    datatype_new ('a, 'b) fn = Fn "'a \<Rightarrow> 'b"
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   321
    datatype_new 'a also_wrong = Also_Wrong (*<*)'a
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   322
    typ (*>*)"('a also_wrong, 'a) fn"
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   323
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   324
text {*
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   325
\noindent
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   326
This is legal:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   327
*}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   328
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   329
    datatype_new 'a ftree = FTLeaf 'a | FTNode "'a \<Rightarrow> 'a ftree"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   330
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   331
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   332
\noindent
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   333
In general, type constructors @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   334
allow recursion on a subset of their type arguments @{text 'a\<^sub>1}, \ldots,
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   335
@{text 'a\<^sub>m}. These type arguments are called \emph{live}; the remaining
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   336
type arguments are called \emph{dead}. In @{typ "'a \<Rightarrow> 'b"} and
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   337
@{typ "('a, 'b) fn"}, the type variable @{typ 'a} is dead and @{typ 'b} is live.
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   338
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   339
Type constructors must be registered as bounded natural functors (BNFs) to have
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   340
live arguments. This is done automatically for datatypes and codatatypes
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   341
introduced by the @{command datatype_new} and @{command codatatype} commands.
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   342
Section~\ref{sec:registering-bounded-natural-functors} explains how to register
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   343
arbitrary type constructors as BNFs.
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   344
*}
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   345
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   346
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   347
subsubsection {* Custom Names and Syntaxes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   348
  \label{sssec:datatype-custom-names-and-syntaxes} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   349
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   350
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   351
The @{command datatype_new} command introduces various constants in addition to
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   352
the constructors. With each datatype are associated set functions, a map
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   353
function, a relator, discriminators, and selectors, all of which can be given
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   354
custom names. In the example below, the traditional names
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   355
@{text set}, @{text map}, @{text list_all2}, @{text null}, @{text hd}, and
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   356
@{text tl} override the default names @{text list_set}, @{text list_map}, @{text
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   357
list_rel}, @{text is_Nil}, @{text un_Cons1}, and @{text un_Cons2}:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   358
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   359
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   360
(*<*)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   361
    no_translations
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   362
      "[x, xs]" == "x # [xs]"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   363
      "[x]" == "x # []"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   364
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   365
    no_notation
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   366
      Nil ("[]") and
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   367
      Cons (infixr "#" 65)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   368
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   369
    hide_type list
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   370
    hide_const Nil Cons hd tl set map list_all2 list_case list_rec
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   371
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   372
    locale dummy_list
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   373
    begin
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   374
(*>*)
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   375
    datatype_new (set: 'a) list (map: map rel: list_all2) =
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   376
      null: Nil (defaults tl: Nil)
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   377
    | Cons (hd: 'a) (tl: "'a list")
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   378
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   379
text {*
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   380
\noindent
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   381
The command introduces a discriminator @{const null} and a pair of selectors
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   382
@{const hd} and @{const tl} characterized as follows:
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   383
%
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   384
\[@{thm list.collapse(1)[of xs, no_vars]}
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   385
  \qquad @{thm list.collapse(2)[of xs, no_vars]}\]
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   386
%
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   387
For two-constructor datatypes, a single discriminator constant suffices. The
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   388
discriminator associated with @{const Cons} is simply
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   389
@{term "\<lambda>xs. \<not> null xs"}.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   390
53553
d4191bf88529 avoid a keyword
blanchet
parents: 53552
diff changeset
   391
The @{text defaults} clause following the @{const Nil} constructor specifies a
d4191bf88529 avoid a keyword
blanchet
parents: 53552
diff changeset
   392
default value for selectors associated with other constructors. Here, it is used
d4191bf88529 avoid a keyword
blanchet
parents: 53552
diff changeset
   393
to ensure that the tail of the empty list is itself (instead of being left
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   394
unspecified).
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   395
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   396
Because @{const Nil} is nullary, it is also possible to use
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   397
@{term "\<lambda>xs. xs = Nil"} as a discriminator. This is specified by
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   398
entering ``@{text "="}'' instead of the identifier @{const null}. Although this
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   399
may look appealing, the mixture of constructors and selectors in the
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   400
characteristic theorems can lead Isabelle's automation to switch between the
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   401
constructor and the destructor view in surprising ways.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   402
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   403
The usual mixfix syntaxes are available for both types and constructors. For
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   404
example:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   405
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   406
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   407
(*<*)
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   408
    end
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   409
(*>*)
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   410
    datatype_new ('a, 'b) prod (infixr "*" 20) = Pair 'a 'b
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   411
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   412
text {* \blankline *}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   413
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   414
    datatype_new (set: 'a) list (map: map rel: list_all2) =
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   415
      null: Nil ("[]")
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   416
    | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   417
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   418
text {*
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   419
\noindent
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   420
Incidentally, this is how the traditional syntaxes can be set up:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   421
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   422
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   423
    syntax "_list" :: "args \<Rightarrow> 'a list" ("[(_)]")
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   424
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   425
text {* \blankline *}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   426
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   427
    translations
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   428
      "[x, xs]" == "x # [xs]"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   429
      "[x]" == "x # []"
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   430
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   431
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   432
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   433
  \label{ssec:datatype-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   434
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   435
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   436
subsubsection {* \keyw{datatype\_new}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   437
  \label{sssec:datatype-new} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   438
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   439
text {*
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   440
Datatype definitions have the following general syntax:
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   441
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   442
@{rail "
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   443
  @@{command_def datatype_new} target? @{syntax dt_options}? \\
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   444
    (@{syntax dt_name} '=' (@{syntax ctor} + '|') + @'and')
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   445
  ;
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   446
  @{syntax_def dt_options}: '(' (('no_discs_sels' | 'rep_compat') + ',') ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   447
"}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   448
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   449
The syntactic quantity \synt{target} can be used to specify a local
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   450
context---e.g., @{text "(in linorder)"}. It is documented in the Isar reference
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   451
manual \cite{isabelle-isar-ref}.
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   452
%
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   453
The optional target is optionally followed by datatype-specific options:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   454
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   455
\begin{itemize}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   456
\setlength{\itemsep}{0pt}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   457
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   458
\item
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   459
The @{text "no_discs_sels"} option indicates that no discriminators or selectors
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   460
should be generated.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   461
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   462
\item
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   463
The @{text "rep_compat"} option indicates that the names generated by the
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   464
package should contain optional (and normally not displayed) ``@{text "new."}''
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   465
components to prevent clashes with a later call to \keyw{rep\_datatype}. See
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   466
Section~\ref{ssec:datatype-compatibility-issues} for details.
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   467
\end{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   468
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   469
The left-hand sides of the datatype equations specify the name of the type to
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   470
define, its type parameters, and additional information:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   471
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   472
@{rail "
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   473
  @{syntax_def dt_name}: @{syntax tyargs}? name @{syntax map_rel}? mixfix?
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   474
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   475
  @{syntax_def tyargs}: typefree | '(' ((name ':')? typefree + ',') ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   476
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   477
  @{syntax_def map_rel}: '(' ((('map' | 'rel') ':' name) +) ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   478
"}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   479
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   480
\noindent
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   481
The syntactic quantity \synt{name} denotes an identifier, \synt{typefree}
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   482
denotes fixed type variable (@{typ 'a}, @{typ 'b}, \ldots), and \synt{mixfix}
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   483
denotes the usual parenthesized mixfix notation. They are documented in the Isar
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   484
reference manual \cite{isabelle-isar-ref}.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   485
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   486
The optional names preceding the type variables allow to override the default
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   487
names of the set functions (@{text t_set1}, \ldots, @{text t_setM}).
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   488
Inside a mutually recursive datatype specification, all defined datatypes must
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   489
specify exactly the same type variables in the same order.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   490
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   491
@{rail "
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   492
  @{syntax_def ctor}: (name ':')? name (@{syntax ctor_arg} * ) \\
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   493
    @{syntax dt_sel_defaults}? mixfix?
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   494
"}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   495
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   496
\medskip
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   497
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   498
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   499
The main constituents of a constructor specification is the name of the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   500
constructor and the list of its argument types. An optional discriminator name
53554
78fe0002024d more (co)data docs
blanchet
parents: 53553
diff changeset
   501
can be supplied at the front to override the default name
78fe0002024d more (co)data docs
blanchet
parents: 53553
diff changeset
   502
(@{text t.is_C\<^sub>j}).
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   503
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   504
@{rail "
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
   505
  @{syntax_def ctor_arg}: type | '(' name ':' type ')'
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   506
"}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   507
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   508
\medskip
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   509
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   510
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   511
In addition to the type of a constructor argument, it is possible to specify a
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   512
name for the corresponding selector to override the default name
53554
78fe0002024d more (co)data docs
blanchet
parents: 53553
diff changeset
   513
(@{text un_C\<^sub>ji}). The same selector names can be reused for several
78fe0002024d more (co)data docs
blanchet
parents: 53553
diff changeset
   514
constructors as long as they share the same type.
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   515
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   516
@{rail "
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   517
  @{syntax_def dt_sel_defaults}: '(' 'defaults' (name ':' term +) ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   518
"}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   519
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   520
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   521
Given a constructor
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   522
@{text "C \<Colon> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<sigma>\<^sub>p \<Rightarrow> \<sigma>"},
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   523
default values can be specified for any selector
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   524
@{text "un_D \<Colon> \<sigma> \<Rightarrow> \<tau>"}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   525
associated with other constructors. The specified default value must be of type
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   526
@{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<sigma>\<^sub>p \<Rightarrow> \<tau>"}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   527
(i.e., it may depends on @{text C}'s arguments).
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   528
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   529
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   530
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   531
subsubsection {* \keyw{datatype\_new\_compat}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   532
  \label{sssec:datatype-new-compat} *}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   533
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   534
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   535
The old datatype package provides some functionality that is not yet replicated
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   536
in the new package:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   537
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   538
\begin{itemize}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   539
\item It is integrated with \keyw{fun} and \keyw{function}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   540
\cite{isabelle-function}, Nitpick \cite{isabelle-nitpick}, Quickcheck,
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   541
and other packages.
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   542
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   543
\item It is extended by various add-ons, notably to produce instances of the
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   544
@{const size} function.
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   545
\end{itemize}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   546
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   547
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   548
New-style datatypes can in most case be registered as old-style datatypes using
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   549
the command
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   550
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   551
@{rail "
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   552
  @@{command_def datatype_new_compat} names
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   553
"}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   554
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   555
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   556
where the \textit{names} argument is simply a space-separated list of type names
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   557
that are mutually recursive. For example:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   558
*}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   559
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   560
    datatype_new_compat even_nat odd_nat
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   561
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   562
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   563
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   564
    thm even_nat_odd_nat.size
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   565
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   566
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   567
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   568
    ML {* Datatype_Data.get_info @{theory} @{type_name even_nat} *}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   569
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   570
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   571
For nested recursive datatypes, all types through which recursion takes place
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   572
must be new-style datatypes. In principle, it should be possible to support
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   573
old-style datatypes as well, but the command does not support this yet.
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   574
*}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   575
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   576
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   577
subsection {* Generated Constants
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   578
  \label{ssec:datatype-generated-constants} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   579
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   580
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   581
Given a datatype @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   582
with $m > 0$ live type variables and $n$ constructors
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   583
@{text "t.C\<^sub>1"}, \ldots, @{text "t.C\<^sub>n"}, the
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   584
following auxiliary constants are introduced:
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   585
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   586
\begin{itemize}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   587
\setlength{\itemsep}{0pt}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   588
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   589
\item \relax{Case combinator}: @{text t_case} (rendered using the familiar
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   590
@{text case}--@{text of} syntax)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   591
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   592
\item \relax{Discriminators}: @{text "t.is_C\<^sub>1"}, \ldots,
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   593
@{text "t.is_C\<^sub>n"}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   594
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   595
\item \relax{Selectors}:
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   596
@{text t.un_C\<^sub>11}$, \ldots, @{text t.un_C\<^sub>1k\<^sub>1}, \\
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   597
\phantom{\relax{Selectors:}} \quad\vdots \\
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   598
\phantom{\relax{Selectors:}} @{text t.un_C\<^sub>n1}$, \ldots, @{text t.un_C\<^sub>nk\<^sub>n}.
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   599
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   600
\item \relax{Set functions} (or \relax{natural transformations}):
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   601
@{text t_set1}, \ldots, @{text t_setm}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   602
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   603
\item \relax{Map function} (or \relax{functorial action}): @{text t_map}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   604
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   605
\item \relax{Relator}: @{text t_rel}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   606
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   607
\item \relax{Iterator}: @{text t_fold}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   608
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   609
\item \relax{Recursor}: @{text t_rec}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   610
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   611
\end{itemize}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   612
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   613
\noindent
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   614
The case combinator, discriminators, and selectors are collectively called
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   615
\emph{destructors}. The prefix ``@{text "t."}'' is an optional component of the
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   616
name and is normally hidden. 
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   617
*}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   618
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   619
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   620
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   621
  \label{ssec:datatype-generated-theorems} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   622
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   623
text {*
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   624
The characteristic theorems generated by @{command datatype_new} are grouped in
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   625
three broad categories:
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   626
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   627
\begin{itemize}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   628
\item The \emph{free constructor theorems} are properties about the constructors
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   629
and destructors that can be derived for any freely generated type. Internally,
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   630
the derivation is performed by @{command wrap_free_constructors}.
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   631
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   632
\item The \emph{functorial theorems} are properties of datatypes related to
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   633
their BNF nature.
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   634
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   635
\item The \emph{inductive theorems} are properties of datatypes related to
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   636
their inductive nature.
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   637
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   638
\end{itemize}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   639
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   640
\noindent
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   641
The full list of named theorems can be obtained as usual by entering the
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   642
command \keyw{print\_theorems} immediately after the datatype definition.
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   643
This list normally excludes low-level theorems that reveal internal
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   644
constructions. To make these accessible, add the line
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   645
*}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   646
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   647
    declare [[bnf_note_all]]
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   648
(*<*)
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   649
    declare [[bnf_note_all = false]]
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   650
(*>*)
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   651
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   652
text {*
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   653
\noindent
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   654
to the top of the theory file.
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   655
*}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   656
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   657
subsubsection {* Free Constructor Theorems
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   658
  \label{sssec:free-constructor-theorems} *}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   659
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   660
(*<*)
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   661
    consts is_Cons :: 'a
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   662
(*>*)
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   663
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   664
text {*
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   665
The first subgroup of properties are concerned with the constructors.
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   666
They are listed below for @{typ "'a list"}:
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   667
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   668
\begin{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   669
\begin{description}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   670
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   671
\item[@{text "t."}\hthm{inject} @{text "[iff, induct_simp]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   672
@{thm list.inject[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   673
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   674
\item[@{text "t."}\hthm{distinct} @{text "[simp, induct_simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   675
@{thm list.distinct(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   676
@{thm list.distinct(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   677
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   678
\item[@{text "t."}\hthm{exhaust} @{text "[cases t, case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   679
@{thm list.exhaust[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   680
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   681
\item[@{text "t."}\hthm{nchotomy}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   682
@{thm list.nchotomy[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   683
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   684
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   685
\end{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   686
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   687
\noindent
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   688
In addition, these nameless theorems are registered as safe elimination rules:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   689
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   690
\begin{indentblock}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   691
\begin{description}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   692
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   693
\item[@{text "t."}\hthm{list.distinct {\upshape[}THEN notE}@{text ", elim!"}\hthm{\upshape]}\rm:] ~ \\
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   694
@{thm list.distinct(1)[THEN notE, elim!, no_vars]} \\
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   695
@{thm list.distinct(2)[THEN notE, elim!, no_vars]}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   696
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   697
\end{description}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   698
\end{indentblock}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   699
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   700
\noindent
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   701
The next subgroup is concerned with the case combinator:
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   702
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   703
\begin{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   704
\begin{description}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   705
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   706
\item[@{text "t."}\hthm{case} @{text "[simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   707
@{thm list.case(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   708
@{thm list.case(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   709
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   710
\item[@{text "t."}\hthm{case\_cong}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   711
@{thm list.case_cong[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   712
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   713
\item[@{text "t."}\hthm{weak\_case\_cong} @{text "[cong]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   714
@{thm list.weak_case_cong[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   715
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   716
\item[@{text "t."}\hthm{split}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   717
@{thm list.split[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   718
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   719
\item[@{text "t."}\hthm{split\_asm}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   720
@{thm list.split_asm[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   721
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   722
\item[@{text "t."}\hthm{splits} = @{text "split split_asm"}]
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   723
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   724
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   725
\end{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   726
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   727
\noindent
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   728
The third and last subgroup revolves around discriminators and selectors:
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   729
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   730
\begin{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   731
\begin{description}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   732
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   733
\item[@{text "t."}\hthm{discs} @{text "[simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   734
@{thm list.discs(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   735
@{thm list.discs(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   736
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   737
\item[@{text "t."}\hthm{sels} @{text "[simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   738
@{thm list.sels(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   739
@{thm list.sels(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   740
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   741
\item[@{text "t."}\hthm{collapse} @{text "[simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   742
@{thm list.collapse(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   743
@{thm list.collapse(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   744
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   745
\item[@{text "t."}\hthm{disc\_exclude}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   746
These properties are missing for @{typ "'a list"} because there is only one
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   747
proper discriminator. Had the datatype been introduced with a second
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   748
discriminator called @{const is_Cons}, they would have read thusly: \\[\jot]
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   749
@{prop "null list \<Longrightarrow> \<not> is_Cons list"} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   750
@{prop "is_Cons list \<Longrightarrow> \<not> null list"}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   751
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   752
\item[@{text "t."}\hthm{disc\_exhaust} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   753
@{thm list.disc_exhaust[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   754
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   755
\item[@{text "t."}\hthm{expand}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   756
@{thm list.expand[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   757
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   758
\item[@{text "t."}\hthm{case\_conv}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   759
@{thm list.case_conv[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   760
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   761
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   762
\end{indentblock}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   763
*}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   764
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   765
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   766
subsubsection {* Functorial Theorems
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   767
  \label{sssec:functorial-theorems} *}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   768
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   769
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   770
The BNF-related theorem are as follows:
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   771
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   772
\begin{indentblock}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   773
\begin{description}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   774
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   775
\item[@{text "t."}\hthm{sets} @{text "[code]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   776
@{thm list.sets(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   777
@{thm list.sets(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   778
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   779
\item[@{text "t."}\hthm{map} @{text "[code]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   780
@{thm list.map(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   781
@{thm list.map(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   782
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   783
\item[@{text "t."}\hthm{rel\_inject} @{text "[code]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   784
@{thm list.rel_inject(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   785
@{thm list.rel_inject(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   786
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   787
\item[@{text "t."}\hthm{rel\_distinct} @{text "[code]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   788
@{thm list.rel_distinct(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   789
@{thm list.rel_distinct(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   790
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   791
\end{description}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   792
\end{indentblock}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   793
*}
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   794
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   795
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   796
subsubsection {* Inductive Theorems
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   797
  \label{sssec:inductive-theorems} *}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   798
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   799
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   800
The inductive theorems are as follows:
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   801
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   802
\begin{indentblock}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   803
\begin{description}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   804
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   805
\item[@{text "t."}\hthm{induct} @{text "[induct t, case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   806
@{thm list.induct[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   807
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   808
\item[@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   809
Given $m > 1$ mutually recursive datatypes, this induction rule can be used to
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   810
prove $m$ properties simultaneously.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   811
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   812
\item[@{text "t."}\hthm{fold} @{text "[code]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   813
@{thm list.fold(1)[no_vars]} \\
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   814
@{thm list.fold(2)[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   815
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   816
\item[@{text "t."}\hthm{rec} @{text "[code]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   817
@{thm list.rec(1)[no_vars]} \\
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   818
@{thm list.rec(2)[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   819
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   820
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   821
\end{indentblock}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   822
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   823
\noindent
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   824
For convenience, @{command datatype_new} also provides the following collection:
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   825
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   826
\begin{indentblock}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   827
\begin{description}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   828
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   829
\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}] ~ \\
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   830
@{text t.rel_distinct} @{text t.sets}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   831
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   832
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   833
\end{indentblock}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   834
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   835
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   836
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   837
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   838
subsection {* Compatibility Issues
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   839
  \label{ssec:datatype-compatibility-issues} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   840
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   841
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   842
%  * main incompabilities between two packages
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   843
%    * differences in theorem names (e.g. singular vs. plural for some props?)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   844
%    * differences in "simps"?
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   845
%    * different recursor/induction in nested case
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   846
%        * discussed in Section~\ref{sec:defining-recursive-functions}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   847
%    * different ML interfaces / extension mechanisms
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   848
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   849
%  * register new datatype as old datatype
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   850
%    * motivation:
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   851
%      * do recursion through new datatype in old datatype
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   852
%        (e.g. useful when porting to the new package one type at a time)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   853
%      * use old primrec
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   854
%      * use fun
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   855
%      * use extensions like size (needed for fun), the AFP order, Quickcheck,
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   856
%        Nitpick
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   857
%      * provide exactly the same theorems with the same names (compatibility)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   858
%    * option 1
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   859
%      * @{text "rep_compat"}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   860
%      * \keyw{rep\_datatype}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   861
%      * has some limitations
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   862
%        * mutually recursive datatypes? (fails with rep_datatype?)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   863
%        * nested datatypes? (fails with datatype_new?)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   864
%    * option 2
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   865
%      * @{command datatype_new_compat}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   866
%      * not fully implemented yet?
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   867
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   868
%  * register old datatype as new datatype
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   869
%    * no clean way yet
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   870
%    * if the goal is to do recursion through old datatypes, can register it as
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   871
%      a BNF (Section XXX)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   872
%    * can also derive destructors etc. using @{command wrap_free_constructors}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   873
%      (Section XXX)
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   874
*}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   875
*)
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   876
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   877
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   878
section {* Defining Recursive Functions
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   879
  \label{sec:defining-recursive-functions} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   880
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   881
text {*
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   882
This describes how to specify recursive functions over datatypes specified using
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   883
@{command datatype_new}. The focus in on the @{command primrec_new} command,
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   884
which supports primitive recursion. More examples can be found in
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   885
\verb|~~/src/HOL/BNF/Examples|.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   886
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   887
%%% TODO: partial_function
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   888
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   889
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   890
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   891
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   892
  \label{ssec:primrec-introductory-examples} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   893
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   894
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   895
subsubsection {* Nonrecursive Types
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   896
  \label{sssec:primrec-nonrecursive-types} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   897
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   898
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   899
Primitive recursion removes one layer of constructors on the left-hand side in
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   900
each equation. For example:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   901
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   902
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   903
    primrec_new bool_of_trool :: "trool \<Rightarrow> bool" where
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   904
      "bool_of_trool Faalse \<longleftrightarrow> False" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   905
      "bool_of_trool Truue \<longleftrightarrow> True"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   906
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   907
text {* \blankline *}
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   908
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   909
    primrec_new the_list :: "'a option \<Rightarrow> 'a list" where
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   910
      "the_list None = []" |
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   911
      "the_list (Some a) = [a]"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   912
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   913
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   914
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   915
    primrec_new the_default :: "'a \<Rightarrow> 'a option \<Rightarrow> 'a" where
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   916
      "the_default d None = d" |
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   917
      "the_default _ (Some a) = a"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   918
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   919
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   920
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   921
    primrec_new mirrror :: "('a, 'b, 'c) triple \<Rightarrow> ('c, 'b, 'a) triple" where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   922
      "mirrror (Triple a b c) = Triple c b a"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   923
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   924
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   925
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   926
The equations can be specified in any order, and it is acceptable to leave out
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   927
some cases, which are then unspecified. Pattern matching on the left-hand side
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   928
is restricted to a single datatype, which must correspond to the same argument
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   929
in all equations.
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   930
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   931
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   932
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   933
subsubsection {* Simple Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   934
  \label{sssec:primrec-simple-recursion} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   935
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   936
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   937
For simple recursive types, recursive calls on a constructor argument are
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   938
allowed on the right-hand side:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   939
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   940
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   941
(*<*)
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   942
    context dummy_tlist
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   943
    begin
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   944
(*>*)
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
   945
    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
   946
      "replicate Zero _ = []" |
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
   947
      "replicate (Suc n) a = a # replicate n a"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   948
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   949
text {* \blankline *}
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   950
53332
c97a05a26dd6 Doc improvements
traytel
parents: 53331
diff changeset
   951
    primrec_new at :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" where
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   952
      "at (a # as) j =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   953
         (case j of
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   954
            Zero \<Rightarrow> a
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   955
          | Suc j' \<Rightarrow> at as j')"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   956
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   957
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   958
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   959
    primrec_new tfold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) tlist \<Rightarrow> 'b" where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   960
      "tfold _ (TNil b) = b" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   961
      "tfold f (TCons a as) = f a (tfold f as)"
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   962
(*<*)
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   963
    end
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   964
(*>*)
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   965
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   966
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   967
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   968
The next example is not primitive recursive, but it can be defined easily using
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   969
@{command fun}. The @{command datatype_new_compat} command is needed to register
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   970
new-style datatypes for use with @{command fun} and @{command function}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   971
(Section~\ref{sssec:datatype-new-compat}):
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   972
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   973
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   974
    datatype_new_compat nat
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   975
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   976
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   977
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   978
    fun at_least_two :: "nat \<Rightarrow> bool" where
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   979
      "at_least_two (Suc (Suc _)) \<longleftrightarrow> True" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   980
      "at_least_two _ \<longleftrightarrow> False"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   981
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   982
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   983
subsubsection {* Mutual Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   984
  \label{sssec:primrec-mutual-recursion} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   985
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   986
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   987
The syntax for mutually recursive functions over mutually recursive datatypes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   988
is straightforward:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   989
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   990
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   991
    primrec_new
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   992
      nat_of_even_nat :: "even_nat \<Rightarrow> nat" and
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   993
      nat_of_odd_nat :: "odd_nat \<Rightarrow> nat"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   994
    where
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   995
      "nat_of_even_nat Even_Zero = Zero" |
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   996
      "nat_of_even_nat (Even_Suc n) = Suc (nat_of_odd_nat n)" |
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   997
      "nat_of_odd_nat (Odd_Suc n) = Suc (nat_of_even_nat n)"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   998
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   999
    primrec_new
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1000
      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
  1001
      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
  1002
      eval\<^sub>f :: "('a \<Rightarrow> int) \<Rightarrow> ('b \<Rightarrow> int) \<Rightarrow> ('a, 'b) fct \<Rightarrow> int"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1003
    where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1004
      "eval\<^sub>e \<gamma> \<xi> (Term t) = eval\<^sub>t \<gamma> \<xi> t" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1005
      "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
  1006
      "eval\<^sub>t \<gamma> \<xi> (Factor f) = eval\<^sub>f \<gamma> \<xi> f" |
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1007
      "eval\<^sub>t \<gamma> \<xi> (Prod f t) = eval\<^sub>f \<gamma> \<xi> f + eval\<^sub>t \<gamma> \<xi> t" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1008
      "eval\<^sub>f \<gamma> _ (Const a) = \<gamma> a" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1009
      "eval\<^sub>f _ \<xi> (Var b) = \<xi> b" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1010
      "eval\<^sub>f \<gamma> \<xi> (Expr e) = eval\<^sub>e \<gamma> \<xi> e"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1011
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1012
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1013
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1014
Mutual recursion is be possible within a single type, using @{command fun}:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1015
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1016
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1017
    fun
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1018
      even :: "nat \<Rightarrow> bool" and
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1019
      odd :: "nat \<Rightarrow> bool"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1020
    where
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1021
      "even Zero = True" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1022
      "even (Suc n) = odd n" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1023
      "odd Zero = False" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1024
      "odd (Suc n) = even n"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1025
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1026
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1027
subsubsection {* Nested Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1028
  \label{sssec:primrec-nested-recursion} *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1029
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1030
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1031
In a departure from the old datatype package, nested recursion is normally
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1032
handled via the map functions of the nesting type constructors. For example,
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1033
recursive calls are lifted to lists using @{const map}:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1034
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1035
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1036
(*<*)
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1037
    datatype_new 'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f 'a "'a tree\<^sub>f\<^sub>f list"
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1038
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1039
    context dummy_tlist
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1040
    begin
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1041
(*>*)
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1042
    primrec_new at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \<Rightarrow> nat list \<Rightarrow> 'a" where
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1043
      "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1044
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1045
            [] \<Rightarrow> a
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1046
          | j # js' \<Rightarrow> at (map (\<lambda>t. at\<^sub>f\<^sub>f t js') ts) j)"
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1047
(*<*)
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1048
    end
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1049
(*>*)
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1050
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1051
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1052
The next example features recursion through the @{text option} type. Although
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1053
@{text option} is not a new-style datatype, it is registered as a BNF with the
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1054
map function @{const option_map}:
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1055
*}
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1056
53335
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1057
(*<*)
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1058
    locale sum_btree_nested
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1059
    begin
53335
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1060
(*>*)
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1061
    primrec_new sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" where
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1062
      "sum_btree (BNode a lt rt) =
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1063
         a + the_default 0 (option_map sum_btree lt) +
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1064
           the_default 0 (option_map sum_btree rt)"
53335
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1065
(*<*)
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1066
    end
585b2fee55e5 fixed bug in primrec_new (allow indirect recursion through constructor arguments other than the first)
panny
parents: 53332
diff changeset
  1067
(*>*)
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1068
53136
98a2c33d5d1b ideas for (co)datatype docs
blanchet
parents: 53028
diff changeset
  1069
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1070
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1071
The same principle applies for arbitrary type constructors through which
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1072
recursion is possible. Notably, the map function for the function type
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1073
(@{text \<Rightarrow>}) is simply composition (@{text "op \<circ>"}):
53136
98a2c33d5d1b ideas for (co)datatype docs
blanchet
parents: 53028
diff changeset
  1074
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1075
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1076
    primrec_new ftree_map :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1077
      "ftree_map f (FTLeaf x) = FTLeaf (f x)" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1078
      "ftree_map f (FTNode g) = FTNode (ftree_map f \<circ> g)"
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1079
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1080
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1081
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1082
(No such function is defined by the package because @{typ 'a} is dead in
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1083
@{typ "'a ftree"}, but we can easily do it ourselves.)
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1084
*}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1085
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1086
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1087
subsubsection {* Nested-as-Mutual Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1088
  \label{sssec:datatype-nested-as-mutual-recursion} *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1089
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1090
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1091
For compatibility with the old package, but also because it is sometimes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1092
convenient in its own right, it is possible to treat nested recursive datatypes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1093
as mutually recursive ones if the recursion takes place though new-style
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1094
datatypes. For example:
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1095
*}
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1096
53331
20440c789759 prove theorem in the right context (that knows about local variables)
traytel
parents: 53330
diff changeset
  1097
    primrec_new
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1098
      at_tree\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \<Rightarrow> nat list \<Rightarrow> 'a" and
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1099
      at_trees\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f list \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> 'a"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1100
    where
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1101
      "at_tree\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1102
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1103
            [] \<Rightarrow> a
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1104
          | j # js' \<Rightarrow> at_trees\<^sub>f\<^sub>f ts j js')" |
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1105
      "at_trees\<^sub>f\<^sub>f (t # ts) j =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1106
         (case j of
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1107
            Zero \<Rightarrow> at_tree\<^sub>f\<^sub>f t
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1108
          | Suc j' \<Rightarrow> at_trees\<^sub>f\<^sub>f ts j')"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1109
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1110
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1111
53331
20440c789759 prove theorem in the right context (that knows about local variables)
traytel
parents: 53330
diff changeset
  1112
    primrec_new
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1113
      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
  1114
      sum_btree_option :: "'a btree option \<Rightarrow> 'a"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1115
    where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1116
      "sum_btree (BNode a lt rt) =
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1117
         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
  1118
      "sum_btree_option None = 0" |
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1119
      "sum_btree_option (Some t) = sum_btree t"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1120
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1121
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1122
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1123
%  * can pretend a nested type is mutually recursive (if purely inductive)
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1124
%  * avoids the higher-order map
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1125
%  * e.g.
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1126
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1127
%  * this can always be avoided;
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1128
%     * e.g. in our previous example, we first mapped the recursive
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1129
%       calls, then we used a generic at function to retrieve the result
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1130
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1131
%  * there's no hard-and-fast rule of when to use one or the other,
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1132
%    just like there's no rule when to use fold and when to use
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1133
%    primrec_new
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1134
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1135
%  * higher-order approach, considering nesting as nesting, is more
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1136
%    compositional -- e.g. we saw how we could reuse an existing polymorphic
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1137
%    at or the_default, whereas @{const at_trees\<^sub>f\<^sub>f} is much more specific
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1138
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1139
%  * but:
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1140
%     * is perhaps less intuitive, because it requires higher-order thinking
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1141
%     * may seem inefficient, and indeed with the code generator the
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1142
%       mutually recursive version might be nicer
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1143
%     * is somewhat indirect -- must apply a map first, then compute a result
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1144
%       (cannot mix)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1145
%     * the auxiliary functions like @{const at_trees\<^sub>f\<^sub>f} are sometimes useful in own right
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1146
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1147
%  * impact on automation unclear
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1148
%
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1149
*}
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1150
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1151
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1152
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1153
  \label{ssec:primrec-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1154
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1155
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1156
subsubsection {* \keyw{primrec\_new}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1157
  \label{sssec:primrec-new} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1158
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1159
text {*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1160
Primitive recursive functions have the following general syntax:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1161
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1162
@{rail "
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1163
  @@{command_def primrec_new} target? fixes \\ @'where'
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1164
    (@{syntax primrec_equation} + '|')
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1165
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1166
  @{syntax_def primrec_equation}: thmdecl? prop
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1167
"}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1168
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1169
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1170
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1171
(*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1172
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1173
  \label{ssec:primrec-generated-theorems} *}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1174
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1175
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1176
%  * synthesized nonrecursive definition
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1177
%  * user specification is rederived from it, exactly as entered
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1178
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1179
%  * induct
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1180
%    * mutualized
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1181
%    * without some needless induction hypotheses if not used
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1182
%  * fold, rec
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1183
%    * mutualized
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1184
*}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1185
*)
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1186
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1187
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1188
subsection {* Recursive Default Values for Selectors
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1189
  \label{ssec:primrec-recursive-default-values-for-selectors} *}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1190
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1191
text {*
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1192
A datatype selector @{text un_D} can have a default value for each constructor
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1193
on which it is not otherwise specified. Occasionally, it is useful to have the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1194
default value be defined recursively. This produces a chicken-and-egg situation
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1195
that may seem unsolvable, because the datatype is not introduced yet at the
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1196
moment when the selectors are introduced. Of course, we can always define the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1197
selectors manually afterward, but we then have to state and prove all the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1198
characteristic theorems ourselves instead of letting the package do it.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1199
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1200
Fortunately, there is a fairly elegant workaround that relies on overloading and
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1201
that avoids the tedium of manual derivations:
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1202
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1203
\begin{enumerate}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1204
\setlength{\itemsep}{0pt}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1205
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1206
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1207
Introduce a fully unspecified constant @{text "un_D\<^sub>0 \<Colon> 'a"} using
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1208
@{keyword consts}.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1209
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1210
\item
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1211
Define the datatype, specifying @{text "un_D\<^sub>0"} as the selector's default
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1212
value.
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1213
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1214
\item
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1215
Define the behavior of @{text "un_D\<^sub>0"} on values of the newly introduced
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1216
datatype using the \keyw{overloading} command.
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1217
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1218
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1219
Derive the desired equation on @{text un_D} from the characteristic equations
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1220
for @{text "un_D\<^sub>0"}.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1221
\end{enumerate}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1222
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1223
\noindent
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1224
The following example illustrates this procedure:
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1225
*}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1226
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1227
    consts termi\<^sub>0 :: 'a
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1228
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1229
text {* \blankline *}
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1230
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
  1231
    datatype_new ('a, 'b) tlist =
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1232
      TNil (termi: 'b) (defaults ttl: TNil)
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
  1233
    | TCons (thd: 'a) (ttl : "('a, 'b) tlist") (defaults termi: "\<lambda>_ xs. termi\<^sub>0 xs")
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1234
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1235
text {* \blankline *}
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1236
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1237
    overloading
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
  1238
      termi\<^sub>0 \<equiv> "termi\<^sub>0 \<Colon> ('a, 'b) tlist \<Rightarrow> 'b"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1239
    begin
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
  1240
    primrec_new termi\<^sub>0 :: "('a, 'b) tlist \<Rightarrow> 'b" where
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1241
      "termi\<^sub>0 (TNil y) = y" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1242
      "termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1243
    end
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1244
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1245
text {* \blankline *}
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1246
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1247
    lemma terminal_TCons[simp]: "termi (TCons x xs) = termi xs"
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1248
    by (cases xs) auto
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1249
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1250
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1251
(*
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1252
subsection {* Compatibility Issues
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1253
  \label{ssec:primrec-compatibility-issues} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1254
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1255
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1256
%  * different induction in nested case
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1257
%    * solution: define nested-as-mutual functions with primrec,
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1258
%      and look at .induct
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1259
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1260
%  * different induction and recursor in nested case
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1261
%    * only matters to low-level users; they can define a dummy function to force
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1262
%      generation of mutualized recursor
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1263
*}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1264
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1265
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1266
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1267
section {* Defining Codatatypes
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1268
  \label{sec:defining-codatatypes} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1269
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1270
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1271
Codatatypes can be specified using the @{command codatatype} command. The
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1272
command is first illustrated through concrete examples featuring different
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1273
flavors of corecursion. More examples can be found in the directory
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1274
\verb|~~/src/HOL/BNF/Examples|. The \emph{Archive of Formal Proofs} also
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1275
includes some useful codatatypes, notably for lazy lists \cite{lochbihler-2010}.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1276
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1277
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1278
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1279
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1280
  \label{ssec:codatatype-introductory-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1281
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1282
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1283
subsubsection {* Simple Corecursion
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1284
  \label{sssec:codatatype-simple-corecursion} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1285
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1286
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1287
Noncorecursive codatatypes coincide with the corresponding datatypes, so
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1288
they have no practical uses. \emph{Corecursive codatatypes} have the same syntax
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1289
as recursive datatypes, except for the command name. For example, here is the
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1290
definition of lazy lists:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1291
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1292
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1293
    codatatype (lset: 'a) llist (map: lmap rel: llist_all2) =
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1294
      lnull: LNil (defaults ltl: LNil)
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1295
    | LCons (lhd: 'a) (ltl: "'a llist")
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1296
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1297
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1298
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1299
Lazy lists can be infinite, such as @{text "LCons 0 (LCons 0 (\<dots>))"} and
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1300
@{text "LCons 0 (LCons 1 (LCons 2 (\<dots>)))"}. Another interesting type that can
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1301
be defined as a codatatype is that of the extended natural numbers:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1302
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1303
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1304
    codatatype enat = EZero | ESuc nat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1305
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1306
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1307
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1308
This type has exactly one infinite element, @{text "ESuc (ESuc (ESuc (\<dots>)))"},
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1309
that represents $\infty$. In addition, it has finite values of the form
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1310
@{text "ESuc (\<dots> (ESuc EZero)\<dots>)"}.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1311
*}
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1312
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1313
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1314
subsubsection {* Mutual Corecursion
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1315
  \label{sssec:codatatype-mutual-corecursion} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1316
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1317
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1318
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1319
The example below introduces a pair of \emph{mutually corecursive} types:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1320
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1321
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1322
    codatatype even_enat = Even_EZero | Even_ESuc odd_enat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1323
    and odd_enat = Odd_ESuc even_enat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1324
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1325
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1326
subsubsection {* Nested Corecursion
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1327
  \label{sssec:codatatype-nested-corecursion} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1328
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1329
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1330
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1331
The next two examples feature \emph{nested corecursion}:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1332
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1333
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1334
    codatatype 'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist"
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1335
    codatatype 'a tree\<^sub>i\<^sub>s = Node\<^sub>i\<^sub>s 'a "'a tree\<^sub>i\<^sub>s fset"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1336
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1337
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1338
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1339
  \label{ssec:codatatype-command-syntax} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1340
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1341
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1342
subsubsection {* \keyw{codatatype}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1343
  \label{sssec:codatatype} *}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1344
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1345
text {*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1346
Definitions of codatatypes have almost exactly the same syntax as for datatypes
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1347
(Section~\ref{ssec:datatype-command-syntax}), with two exceptions: The command
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1348
is called @{command codatatype}. The @{text "no_discs_sels"} option is not
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1349
available, because destructors are a crucial notion for codatatypes.
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1350
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1351
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1352
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1353
subsection {* Generated Constants
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1354
  \label{ssec:codatatype-generated-constants} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1355
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1356
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1357
Given a codatatype @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1358
with $m > 0$ live type variables and $n$ constructors @{text "t.C\<^sub>1"},
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1359
\ldots, @{text "t.C\<^sub>n"}, the same auxiliary constants are generated as for
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1360
datatypes (Section~\ref{ssec:datatype-generated-constants}), except that the
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1361
iterator and the recursor are replaced by dual concepts:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1362
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1363
\begin{itemize}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1364
\setlength{\itemsep}{0pt}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1365
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1366
\item \relax{Coiterator}: @{text t_unfold}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1367
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1368
\item \relax{Corecursor}: @{text t_corec}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1369
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1370
\end{itemize}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1371
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1372
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1373
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1374
subsection {* Generated Theorems
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1375
  \label{ssec:codatatype-generated-theorems} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1376
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1377
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1378
The characteristic theorems generated by @{command codatatype} are grouped in
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1379
three broad categories:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1380
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1381
\begin{itemize}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1382
\item The \emph{free constructor theorems} are properties about the constructors
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1383
and destructors that can be derived for any freely generated type.
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1384
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1385
\item The \emph{functorial theorems} are properties of datatypes related to
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1386
their BNF nature.
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1387
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1388
\item The \emph{coinductive theorems} are properties of datatypes related to
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1389
their coinductive nature.
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1390
\end{itemize}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1391
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1392
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1393
The first two categories are exactly as for datatypes and are described in
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
  1394
Sections
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
  1395
\ref{sssec:free-constructor-theorems}~and~\ref{sssec:functorial-theorems}.
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1396
*}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1397
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1398
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1399
subsubsection {* Coinductive Theorems
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1400
  \label{sssec:coinductive-theorems} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1401
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1402
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1403
The coinductive theorems are as follows:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1404
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1405
%          [(coinductN, map single coinduct_thms,
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1406
%            fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1407
%           (corecN, corec_thmss, K coiter_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1408
%           (disc_corecN, disc_corec_thmss, K disc_coiter_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1409
%           (disc_corec_iffN, disc_corec_iff_thmss, K disc_coiter_iff_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1410
%           (disc_unfoldN, disc_unfold_thmss, K disc_coiter_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1411
%           (disc_unfold_iffN, disc_unfold_iff_thmss, K disc_coiter_iff_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1412
%           (sel_corecN, sel_corec_thmss, K sel_coiter_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1413
%           (sel_unfoldN, sel_unfold_thmss, K sel_coiter_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1414
%           (simpsN, simp_thmss, K []),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1415
%           (strong_coinductN, map single strong_coinduct_thms, K coinduct_attrs),
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1416
%           (unfoldN, unfold_thmss, K coiter_attrs)]
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1417
%          |> massage_multi_notes;
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1418
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1419
\begin{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1420
\begin{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1421
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
  1422
\item[@{text "t."}\hthm{coinduct} @{text "[induct t, case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1423
@{thm llist.coinduct[no_vars]}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1424
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
  1425
\item[@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{coinduct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1426
Given $m > 1$ mutually recursive datatypes, this induction rule can be used to
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1427
prove $m$ properties simultaneously.
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1428
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
  1429
\item[@{text "t."}\hthm{unfold} @{text "[code]"}\rm:] ~ \\
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1430
@{thm llist.unfold(1)[no_vars]} \\
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1431
@{thm llist.unfold(2)[no_vars]}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1432
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
  1433
\item[@{text "t."}\hthm{corec} @{text "[code]"}\rm:] ~ \\
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1434
@{thm llist.corec(1)[no_vars]} \\
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1435
@{thm llist.corec(2)[no_vars]}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1436
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1437
\end{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1438
\end{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1439
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1440
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1441
For convenience, @{command codatatype} also provides the following collection:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1442
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1443
\begin{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1444
\begin{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1445
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1446
\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}] ~ \\
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1447
@{text t.rel_distinct} @{text t.sets}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1448
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1449
\end{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1450
\end{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1451
*}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1452
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1453
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1454
section {* Defining Corecursive Functions
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1455
  \label{sec:defining-corecursive-functions} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1456
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1457
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1458
Corecursive functions can be specified using the @{command primcorec} command.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1459
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1460
%%% TODO: partial_function? E.g. for defining tail recursive function on lazy
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1461
%%% lists (cf. terminal0 in TLList.thy)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1462
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1463
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1464
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1465
(*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1466
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1467
  \label{ssec:primcorec-introductory-examples} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1468
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1469
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1470
More examples in \verb|~~/src/HOL/BNF/Examples|.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1471
*}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1472
*)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1473
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1474
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1475
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1476
  \label{ssec:primcorec-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1477
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1478
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1479
subsubsection {* \keyw{primcorec}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1480
  \label{sssec:primcorec} *}
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1481
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1482
text {*
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1483
Primitive corecursive definitions have the following general syntax:
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1484
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1485
@{rail "
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1486
  @@{command_def primcorec} target? fixes \\ @'where'
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1487
    (@{syntax primcorec_formula} + '|')
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1488
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1489
  @{syntax_def primcorec_formula}: thmdecl? prop (@'of' (term * ))?
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1490
"}
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1491
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1492
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1493
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1494
(*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1495
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1496
  \label{ssec:primcorec-generated-theorems} *}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1497
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1498
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1499
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1500
(*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1501
subsection {* Recursive Default Values for Selectors
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1502
  \label{ssec:primcorec-recursive-default-values-for-selectors} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1503
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1504
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1505
partial_function to the rescue
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1506
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1507
*)
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1508
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1509
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1510
section {* Registering Bounded Natural Functors
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1511
  \label{sec:registering-bounded-natural-functors} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1512
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1513
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1514
The (co)datatype package can be set up to allow nested recursion through custom
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1515
well-behaved type constructors. The key concept is that of a bounded natural
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1516
functor (BNF).
52829
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1517
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1518
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1519
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1520
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1521
(*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1522
subsection {* Introductory Example
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1523
  \label{ssec:bnf-introductory-example} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1524
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1525
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1526
More examples in \verb|~~/src/HOL/BNF/Basic_BNFs.thy| and
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1527
\verb|~~/src/HOL/BNF/More_BNFs.thy|.
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1528
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1529
%Mention distinction between live and dead type arguments;
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1530
%  * and existence of map, set for those
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1531
%mention =>.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1532
*}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1533
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1534
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1535
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1536
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1537
  \label{ssec:bnf-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1538
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1539
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1540
subsubsection {* \keyw{bnf}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1541
  \label{sssec:bnf} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1542
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1543
text {*
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1544
@{rail "
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1545
  @@{command_def bnf} target? (name ':')? term \\
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1546
    term_list term term_list term?
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1547
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1548
  X_list: '[' (X + ',') ']'
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1549
"}
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1550
*}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1551
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1552
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1553
subsubsection {* \keyw{print\_bnfs}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1554
  \label{sssec:print-bnfs} *}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1555
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1556
text {*
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1557
@{command print_bnfs}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1558
*}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1559
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1560
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1561
section {* Deriving Destructors and Theorems for Free Constructors
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1562
  \label{sec:deriving-destructors-and-theorems-for-free-constructors} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1563
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1564
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1565
The derivation of convenience theorems for types equipped with free constructors,
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1566
as performed internally by @{command datatype_new} and @{command codatatype},
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1567
is available as a stand-alone command called @{command wrap_free_constructors}.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1568
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1569
%  * need for this is rare but may arise if you want e.g. to add destructors to
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1570
%    a type not introduced by ...
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1571
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1572
%  * also useful for compatibility with old package, e.g. add destructors to
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1573
%    old \keyw{datatype}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1574
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1575
%  * @{command wrap_free_constructors}
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1576
%    * @{text "no_discs_sels"}, @{text "rep_compat"}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1577
%    * hack to have both co and nonco view via locale (cf. ext nats)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1578
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1579
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1580
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1581
(*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1582
subsection {* Introductory Example
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1583
  \label{ssec:ctors-introductory-example} *}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1584
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1585
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1586
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1587
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1588
  \label{ssec:ctors-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1589
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1590
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1591
subsubsection {* \keyw{wrap\_free\_constructors}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1592
    \label{sssec:wrap-free-constructors} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1593
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1594
text {*
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1595
Free constructor wrapping has the following general syntax:
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1596
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1597
@{rail "
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1598
  @@{command_def wrap_free_constructors} target? @{syntax dt_options} \\
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1599
    term_list name @{syntax fc_discs_sels}?
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1600
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1601
  @{syntax_def fc_discs_sels}: name_list (name_list_list name_term_list_list? )?
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1602
  ;
53534
de2027f9aff3 more (co)datatype documentation
blanchet
parents: 53491
diff changeset
  1603
  @{syntax_def name_term}: (name ':' term)
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1604
"}
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1605
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1606
% options: no_discs_sels rep_compat
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1607
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1608
% X_list is as for BNF
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1609
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
  1610
Section~\ref{ssec:datatype-generated-theorems} lists the generated theorems.
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1611
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1612
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1613
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1614
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1615
section {* Standard ML Interface
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1616
  \label{sec:standard-ml-interface} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1617
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1618
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1619
The package's programmatic interface.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1620
*}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1621
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1622
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1623
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1624
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1625
section {* Interoperability
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1626
  \label{sec:interoperability} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1627
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1628
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1629
The package's interaction with other Isabelle packages and tools, such as the
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1630
code generator and the counterexample generators.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1631
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1632
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1633
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1634
subsection {* Transfer and Lifting
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1635
  \label{ssec:transfer-and-lifting} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1636
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1637
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1638
subsection {* Code Generator
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1639
  \label{ssec:code-generator} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1640
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1641
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1642
subsection {* Quickcheck
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1643
  \label{ssec:quickcheck} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1644
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1645
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1646
subsection {* Nitpick
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1647
  \label{ssec:nitpick} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1648
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1649
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1650
subsection {* Nominal Isabelle
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1651
  \label{ssec:nominal-isabelle} *}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1652
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1653
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1654
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1655
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1656
section {* Known Bugs and Limitations
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1657
  \label{sec:known-bugs-and-limitations} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1658
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1659
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1660
Known open issues of the package.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1661
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1662
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1663
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1664
%* primcorec is unfinished
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1665
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1666
%* slow n-ary mutual (co)datatype, avoid as much as possible (e.g. using nesting)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1667
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1668
%* issues with HOL-Proofs?
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1669
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1670
%* partial documentation
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1671
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1672
%* no way to register "sum" and "prod" as (co)datatypes to enable N2M reduction for them
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1673
%  (for @{command datatype_new_compat} and prim(co)rec)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1674
%
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1675
%    * a fortiori, no way to register same type as both data- and codatatype
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1676
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1677
%* no recursion through unused arguments (unlike with the old package)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1678
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1679
%* in a locale, cannot use locally fixed types (because of limitation in typedef)?
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1680
%
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1681
% *names of variables suboptimal
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1682
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1683
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1684
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1685
section {* Acknowledgments
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1686
  \label{sec:acknowledgments} *}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1687
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1688
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1689
Tobias Nipkow and Makarius Wenzel have encouraged us to implement the new
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1690
(co)datatype package. Andreas Lochbihler provided lots of comments on earlier
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1691
versions of the package, especially for the coinductive part. Brian Huffman
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1692
suggested major simplifications to the internal constructions, much of which has
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1693
yet to be implemented. Florian Haftmann and Christian Urban provided general
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1694
advice advice on Isabelle and package writing. Stefan Milius and Lutz Schr\"oder
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1695
suggested an elegant proof to eliminate one of the BNF assumptions.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1696
*}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1697
*)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1698
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1699
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1700
end