src/Doc/Datatypes/Datatypes.thy
author blanchet
Tue, 13 Aug 2013 19:57:57 +0200
changeset 53018 11ebef554439
parent 52969 f2df0730f8ac
child 53025 c820c9e9e8f4
permissions -rw-r--r--
added rail diagram
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
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     3
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     4
Tutorial for (co)datatype definitions with the new package.
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     5
*)
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
theory Datatypes
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
     8
imports Setup
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
     9
keywords
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    10
  "primrec_new" :: thy_decl and
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    11
  "primcorec" :: thy_decl
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
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    14
(*<*)
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    15
(* FIXME: Evil setup until "primrec_new" and "primcorec" are in place. *)
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    16
ML_command {*
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    17
fun add_dummy_cmd _ _ lthy = lthy;
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    18
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    19
val _ = Outer_Syntax.local_theory @{command_spec "primrec_new"} ""
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    20
  (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    21
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    22
val _ = Outer_Syntax.local_theory @{command_spec "primcorec"} ""
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    23
  (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    24
*}
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    25
(*>*)
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    26
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    27
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    28
section {* Introduction
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    29
  \label{sec:introduction} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    30
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    31
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    32
The 2013 edition of Isabelle introduced new definitional package for datatypes
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    33
and codatatypes. The datatype support is similar to that provided by the earlier
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    34
package due to Berghofer and Wenzel \cite{Berghofer-Wenzel:1999:TPHOL},
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    35
documented in the Isar reference manual \cite{isabelle-isar-ref};
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    36
indeed, replacing the keyword @{command datatype} by @{command datatype_new} is
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    37
usually all that is needed to port existing theories to use the new package.
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    38
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    39
Perhaps the main advantage of the new package is that it supports recursion
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    40
through a large class of non-datatypes, comprising finite sets:
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    41
*}
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    42
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    43
    datatype_new 'a treeFS = NodeFS 'a "'a treeFS fset"
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    44
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    45
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    46
\noindent
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    47
Another strong point is that the package supports local definitions:
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    48
*}
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    49
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    50
    context linorder
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    51
    begin
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    52
    datatype_new flag = Less | Eq | Greater
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    53
    end
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    54
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    55
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    56
\noindent
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    57
The package also provides some convenience, notably automatically generated
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    58
destructors (discriminators and selectors).
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    59
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    60
In addition to plain inductive datatypes, the package supports coinductive
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    61
datatypes, or \emph{codatatypes}, which may have infinite values. For example,
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    62
the following command introduces the type of lazy lists:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    63
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    64
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    65
    codatatype 'a llist = LNil | LCons 'a "'a llist"
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    66
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    67
text {*
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    68
\noindent
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    69
Mixed inductive--coinductive recursion is possible via nesting. Compare the
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    70
following four examples:
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    71
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    72
%%% TODO: Avoid 0
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    73
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    74
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    75
    datatype_new 'a treeFF0 = NodeFF 'a "'a treeFF0 list"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    76
    datatype_new 'a treeFI0 = NodeFI 'a "'a treeFI0 llist"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    77
    codatatype 'a treeIF0 = NodeIF 'a "'a treeIF0 list"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    78
    codatatype 'a treeII0 = NodeII 'a "'a treeII0 llist"
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    79
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    80
text {*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    81
The first two tree types allow only finite branches, whereas the last two allow
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    82
infinite branches. Orthogonally, the nodes in the first and third types have
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    83
finite branching, whereas those of the second and fourth may have infinitely
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
    84
many direct subtrees.
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    85
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    86
To use the package, it is necessary to import the @{theory BNF} theory, which
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
    87
can be precompiled as the \textit{HOL-BNF} image. The following commands show
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
    88
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
    89
without launching jEdit:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    90
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    91
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    92
text {*
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    93
\noindent
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
    94
\ \ \ \ \texttt{isabelle jedit -l HOL-BNF} \\
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
    95
\noindent
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    96
\hbox{}\ \ \ \ \texttt{isabelle build -b HOL-BNF}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    97
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    98
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    99
text {*
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   100
The package, like its predecessor, fully adheres to the LCF philosophy
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   101
\cite{mgordon79}: The characteristic theorems associated with the specified
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   102
(co)datatypes are derived rather than introduced axiomatically.%
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   103
\footnote{If the \textit{quick\_and\_dirty} option is enabled, some of the
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   104
internal constructions and most of the internal proof obligations are skipped.}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   105
The package's metatheory is described in a pair of papers
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   106
\cite{traytel-et-al-2012,blanchette-et-al-wit}.
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   107
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   108
This tutorial is organized as follows:
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   109
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   110
\begin{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   111
\setlength{\itemsep}{0pt}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   112
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   113
\item Section \ref{sec:defining-datatypes}, ``Defining Datatypes,''
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   114
describes how to specify datatypes using the @{command datatype_new} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   115
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   116
\item Section \ref{sec:defining-recursive-functions}, ``Defining Recursive
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   117
Functions,'' describes how to specify recursive functions using
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   118
\keyw{primrec\_new}, @{command fun}, and @{command function}.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   119
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   120
\item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,''
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   121
describes how to specify codatatypes using the @{command codatatype} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   122
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   123
\item Section \ref{sec:defining-corecursive-functions}, ``Defining Corecursive
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   124
Functions,'' describes how to specify corecursive functions using the
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   125
\keyw{primcorec} command.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   126
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   127
\item Section \ref{sec:registering-bounded-natural-functors}, ``Registering
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   128
Bounded Natural Functors,'' explains how to set up the (co)datatype package to
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   129
allow nested recursion through custom well-behaved type constructors.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   130
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   131
\item Section \ref{sec:generating-free-constructor-theorems}, ``Generating Free
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   132
Constructor Theorems,'' explains how to derive convenience theorems for free
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   133
constructors, as performed internally by @{command datatype_new} and
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   134
@{command codatatype}.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   135
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   136
\item Section \ref{sec:standard-ml-interface}, ``Standard ML Interface,''
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   137
describes the package's programmatic interface.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   138
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   139
\item Section \ref{sec:interoperability}, ``Interoperability,''
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   140
is concerned with the packages' interaction with other Isabelle packages and
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   141
tools, such as the code generator and the counterexample generators.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   142
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   143
\item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   144
Limitations,'' concludes with known open issues at the time of writing.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   145
\end{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   146
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   147
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   148
\newbox\boxA
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   149
\setbox\boxA=\hbox{\texttt{nospam}}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   150
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   151
\newcommand\authoremaili{\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@\allowbreak
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   152
in.\allowbreak tum.\allowbreak de}}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   153
\newcommand\authoremailii{\texttt{pope{\color{white}nospam}\kern-\wd\boxA{}scua@\allowbreak
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   154
in.\allowbreak tum.\allowbreak de}}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   155
\newcommand\authoremailiii{\texttt{tray{\color{white}nospam}\kern-\wd\boxA{}tel@\allowbreak
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
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   158
The command @{command datatype_new} is expected to displace @{command datatype}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   159
in a future release. Authors of new theories are encouraged to use
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   160
@{command datatype_new}, and maintainers of older theories may want to consider
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   161
upgrading in the coming months.
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
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   164
directed to the authors at \authoremaili, \authoremailii, and \authoremailiii.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   165
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   166
\begin{framed}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   167
\noindent
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   168
\textbf{Warning:} This tutorial is under heavy construction. Please apologise
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   169
for its appearance and come back in a few months. If you have ideas regarding
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   170
material that should be included, please let the authors know.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   171
\end{framed}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   172
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   173
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   174
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   175
section {* Defining Datatypes
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   176
  \label{sec:defining-datatypes} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   177
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   178
text {*
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   179
This section describes how to specify datatypes using the @{command datatype_new}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   180
command. The command is first illustrated through concrete examples featuring
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   181
different flavors of recursion. More examples can be found in the directory
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   182
\verb|~~/src/HOL/BNF/Examples|.
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   183
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   184
  * libraries include many useful datatypes, e.g. list, option, etc., as well
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   185
    as operations on these;
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   186
    see e.g. ``What's in Main'' \cite{xxx}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   187
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   188
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   189
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   190
subsection {* Examples
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   191
  \label{ssec:datatype-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   192
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   193
subsubsection {* Nonrecursive Types *}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   194
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   195
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   196
enumeration type:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   197
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   198
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   199
    datatype_new trool = Truue | Faalse | Perhaaps
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   200
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   201
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   202
Haskell-style option type:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   203
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   204
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   205
    datatype_new 'a maybe = Nothing | Just 'a
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   206
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   207
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   208
triple:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   209
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   210
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   211
    datatype_new ('a, 'b, 'c) triple = Triple 'a 'b 'c
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   212
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   213
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   214
subsubsection {* Simple Recursion *}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   215
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   216
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   217
simplest recursive type: natural numbers
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   218
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   219
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   220
    datatype_new nat = Zero | Suc nat
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   221
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   222
text {*
52868
cca9e958da5c export ML function (for primcorec)
blanchet
parents: 52843
diff changeset
   223
Setup to be able to write @{text 0} instead of @{const Zero}:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   224
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   225
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   226
    instantiation nat :: zero
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   227
    begin
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   228
    definition "0 = Zero"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   229
    instance ..
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   230
    end
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   231
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   232
text {*
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   233
lists were shown in the introduction
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   234
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   235
terminated lists are a variant:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   236
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   237
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   238
    datatype_new ('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist"
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   239
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   240
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   241
On the right-hand side of the equal sign, the usual Isabelle conventions apply:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   242
Nonatomic types must be enclosed in double quotes.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   243
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   244
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   245
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   246
subsubsection {* Mutual Recursion *}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   247
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   248
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   249
Mutual recursion = Define several types simultaneously, referring to each other.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   250
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   251
Simple example: distinction between even and odd natural numbers:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   252
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   253
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   254
    datatype_new enat = EZero | ESuc onat
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   255
    and onat = OSuc enat
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   256
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   257
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   258
More complex, and more realistic, example:
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   259
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   260
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   261
    datatype_new ('a, 'b) exp =
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   262
      Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   263
    and ('a, 'b) trm =
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   264
      Factor "('a, 'b) fct" | Prod "('a, 'b) fct" "('a, 'b) trm"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   265
    and ('a, 'b) fct =
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   266
      Const 'a | Var 'b | Expr "('a, 'b) exp"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   267
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   268
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   269
subsubsection {* Nested Recursion *}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   270
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   271
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   272
Nested recursion = Have recursion through a type constructor.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   273
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   274
The introduction showed some examples of trees with nesting through lists.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   275
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   276
More complex example, which reuses our maybe:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   277
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   278
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   279
    datatype_new 'a btree =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   280
      BNode 'a "'a btree maybe" "'a btree maybe"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   281
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   282
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   283
Recursion may not be arbitrary; e.g. impossible to define
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   284
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   285
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   286
    datatype_new 'a evil = Evil (*<*)'a
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   287
    typ (*>*)"'a evil \<Rightarrow> 'a evil"
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   288
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   289
text {*
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   290
Issue: => allows recursion only on its right-hand side.
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   291
This issue is inherited by polymorphic datatypes (and codatatypes)
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   292
defined in terms of =>.
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   293
In general, type constructors "('a1, ..., 'an) k" allow recursion on a subset
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   294
of their type arguments.
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   295
*}
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   296
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   297
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   298
subsubsection {* Auxiliary Constants and Syntaxes *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   299
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   300
text {*
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   301
The @{command datatype_new} command introduces various constants in addition to the
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   302
constructors. Given a type @{text "('a1,\<dots>,'aM) t"} with constructors
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   303
@{text "t.C\<^sub>1"}, \ldots, @{text "t.C\<^sub>m"}, the following auxiliary
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   304
constants are introduced (among others):
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   305
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   306
\begin{itemize}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   307
\setlength{\itemsep}{0pt}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   308
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   309
\item \emph{Set functions} (\emph{natural transformations}):
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   310
@{text t_set1}, \ldots, @{text t_setM}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   311
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   312
\item \emph{Map function} (\emph{functorial action}): @{text t_map}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   313
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   314
\item \emph{Relator}: @{text t_rel}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   315
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   316
\item \emph{Iterator}: @{text t_fold}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   317
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   318
\item \emph{Recursor}: @{text t_rec}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   319
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   320
\item \emph{Discriminators}: @{text "t.is_C\<^sub>1"}, \ldots,
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   321
@{text "t.is_C\<^sub>m"}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   322
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   323
\item \emph{Selectors}:
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   324
@{text t.un_C}$_{11}$, \ldots, @{text t.un_C}$_{1n_1}$, \ldots,
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   325
@{text t.un_C}$_{m1}$, \ldots, @{text t.un_C}$_{mn_m}$
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   326
\end{itemize}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   327
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   328
The discriminators and selectors are collectively called \emph{destructors}. The
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   329
@{text "t."} prefix is an optional component of the name and is normally hidden.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   330
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   331
The set functions, map function, relator, discriminators, and selectors can be
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   332
given custom names, as in the example below:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   333
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   334
%%% FIXME: get rid of 0 below
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   335
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   336
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   337
(*<*)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   338
    no_translations
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   339
      "[x, xs]" == "x # [xs]"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   340
      "[x]" == "x # []"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   341
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   342
    no_notation
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   343
      Nil ("[]") and
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   344
      Cons (infixr "#" 65)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   345
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   346
    hide_const Nil Cons hd tl map
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   347
(*>*)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   348
    datatype_new (set0: 'a) list0 (map: map0 rel: list0_all2) =
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   349
      null: Nil (defaults tl: Nil)
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   350
    | Cons (hd: 'a) (tl: "'a list0")
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   351
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   352
text {*
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   353
\noindent
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   354
The command introduces a discriminator @{const null} and a pair of selectors
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   355
@{const hd} and @{const tl} characterized as follows:
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   356
%
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   357
\[@{thm list0.collapse(1)[of xs, no_vars]}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   358
  \qquad @{thm list0.collapse(2)[of xs, no_vars]}\]
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   359
%
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   360
For two-constructor datatypes, a single discriminator constant suffices. The
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   361
discriminator associated with @{const Cons} is simply @{text "\<not> null"}.
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   362
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   363
The \keyw{defaults} keyword following the @{const Nil} constructor specifies a
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   364
default value for selectors associated with other constructors. Here, it is
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   365
used to specify that the tail of the empty list is the empty list (instead of
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   366
being unspecified).
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   367
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   368
Because @{const Nil} is a nullary constructor, it is also possible to use @{text
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   369
"= Nil"} as a discriminator. This is specified by specifying @{text "="} instead
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   370
of the identifier @{const null} in the declaration above. Although this may look
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   371
appealing, the mixture of constructors and selectors in the resulting
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   372
characteristic theorems can lead Isabelle's automation to switch between the
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   373
constructor and the destructor view in surprising ways.
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   374
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   375
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   376
text {*
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   377
The usual mixfix syntaxes are available for both types and constructors. For example:
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   378
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   379
%%% FIXME: remove trailing underscore and use locale trick instead once this is
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   380
%%% supported.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   381
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   382
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   383
    datatype_new ('a, 'b) prod (infixr "*" 20) =
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   384
      Pair 'a 'b
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   385
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   386
(*<*)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   387
    hide_const Nil Cons hd tl
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   388
(*>*)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   389
    datatype_new (set: 'a) list (map: map rel: list_all2) =
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   390
      null: Nil ("[]")
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   391
    | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   392
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   393
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   394
Incidentally, this is how the traditional syntaxes are set up in @{theory List}:
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   395
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   396
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   397
    syntax "_list" :: "args \<Rightarrow> 'a list" ("[(_)]")
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   398
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   399
    translations
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   400
      "[x, xs]" == "x # [xs]"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   401
      "[x]" == "x # []"
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   402
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   403
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   404
subsection {* Syntax
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   405
  \label{ssec:datatype-syntax} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   406
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   407
text {*
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   408
Datatype definitions have the following general syntax:
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   409
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   410
@{rail "
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   411
  @@{command datatype_new} @{syntax target}? @{syntax dt_options}? \\
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   412
    (@{syntax dt_name} '=' (@{syntax ctor} + '|') + @'and')
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   413
  ;
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52868
diff changeset
   414
  @{syntax_def dt_options}: '(' ((@'no_discs_sels' | @'rep_compat') + ',') ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   415
"}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   416
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   417
The syntactic quantity @{syntax target} can be used to specify a local context
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   418
(e.g., @{text "(in linorder)"}). It is documented in the Isar reference manual
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   419
\cite{isabelle-isar-ref}.
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   420
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   421
The optional target is followed by optional options:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   422
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   423
\begin{itemize}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   424
\setlength{\itemsep}{0pt}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   425
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   426
\item
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52868
diff changeset
   427
The \keyw{no\_discs\_sels} option indicates that no destructors (i.e.,
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   428
discriminators and selectors) should be generated.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   429
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   430
\item
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   431
The \keyw{rep\_compat} option indicates that the names generated by the
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   432
package should contain optional (and normally not displayed) @{text "new."}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   433
components to prevent clashes with a later call to @{command rep_datatype}. See
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   434
Section~\ref{ssec:datatype-compatibility-issues} for details.
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   435
\end{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   436
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   437
The left-hand sides of the datatype equations specify the name of the type to
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   438
define, its type parameters, and optional additional information:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   439
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   440
@{rail "
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   441
  @{syntax_def dt_name}: @{syntax tyargs}? @{syntax name}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   442
    @{syntax map_rel}? @{syntax mixfix}?
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   443
  ;
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   444
  @{syntax_def tyargs}: @{syntax typefree} | '(' ((@{syntax name} ':')? @{syntax typefree} + ',') ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   445
  ;
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   446
  @{syntax_def map_rel}: '(' ((('map' | 'rel') ':' @{syntax name}) +) ')'
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   447
"}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   448
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   449
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   450
The syntactic quantity @{syntax name} denotes an identifier, @{syntax typefree}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   451
denotes fixed type variable (@{typ 'a}, @{typ 'b}, \ldots), and @{syntax
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   452
mixfix} denotes the usual parenthesized mixfix notation. They are documented in
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   453
the Isar reference manual \cite{isabelle-isar-ref}.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   454
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   455
The optional names preceding the type variables allow to override the default
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   456
names of the set functions (@{text t_set1}, \ldots, @{text t_setM}).
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   457
Inside a mutually recursive datatype specification, all defined datatypes must
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   458
specify exactly the same type variables in the same order.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   459
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   460
@{rail "
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   461
  @{syntax_def ctor}: (@{syntax name} ':')? @{syntax name} (@{syntax ctor_arg} * ) \\
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   462
    @{syntax dt_sel_defaults}? @{syntax mixfix}?
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   463
"}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   464
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   465
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   466
The main constituents of a constructor specification is the name of the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   467
constructor and the list of its argument types. An optional discriminator name
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   468
can be supplied at the front to override the default name
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   469
(@{text t.un_C}$_{ij}$).
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   470
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   471
@{rail "
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   472
  @{syntax_def ctor_arg}: @{syntax type} | '(' @{syntax name} ':' @{syntax type} ')'
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   473
"}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   474
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   475
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   476
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
   477
name for the corresponding selector to override the default name
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   478
(@{text t.un_C}$_{ij}$). The same selector names can be reused for several
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   479
constructors as long as they have the same type.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   480
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   481
@{rail "
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   482
  @{syntax_def dt_sel_defaults}: '(' @'defaults' (@{syntax name} ':' @{syntax term} +) ')'
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   483
"}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   484
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   485
\noindent
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   486
Given a constructor
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   487
@{text "C \<Colon> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<sigma>\<^sub>p \<Rightarrow> \<sigma>"},
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   488
default values can be specified for any selector
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   489
@{text "un_D \<Colon> \<sigma> \<Rightarrow> \<tau>"}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   490
associated with other constructors. The specified default value must have type
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   491
@{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<sigma>\<^sub>p \<Rightarrow> \<tau>"}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   492
(i.e., it may dependend on @{text C}'s arguments).
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   493
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   494
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   495
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   496
  \label{ssec:datatype-generated-theorems} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   497
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   498
text {*
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   499
  * free ctor theorems
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   500
    * case syntax
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   501
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   502
  * per-type theorems
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   503
    * sets, map, rel
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   504
    * induct, fold, rec
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   505
    * simps
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   506
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   507
  * multi-type (``common'') theorems
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   508
    * induct
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   509
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   510
  * mention what is registered with which attribute
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   511
    * and also nameless safes
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   512
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   513
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   514
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   515
subsection {* Compatibility Issues
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   516
  \label{ssec:datatype-compatibility-issues} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   517
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   518
text {*
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   519
  * main incompabilities between two packages
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   520
    * differences in theorem names (e.g. singular vs. plural for some props?)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   521
    * differences in "simps"?
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   522
    * different recursor/induction in nested case
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   523
        * discussed in Section~\ref{sec:defining-recursive-functions}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   524
    * different ML interfaces / extension mechanisms
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   525
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   526
  * register new datatype as old datatype
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   527
    * motivation:
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   528
      * do recursion through new datatype in old datatype
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   529
        (e.g. useful when porting to the new package one type at a time)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   530
      * use old primrec
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   531
      * use fun
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   532
      * use extensions like size (needed for fun), the AFP order, Quickcheck,
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   533
        Nitpick
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   534
      * provide exactly the same theorems with the same names (compatibility)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   535
    * option 1
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   536
      * \keyw{rep\_compat}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   537
      * \keyw{rep\_datatype}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   538
      * has some limitations
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   539
        * mutually recursive datatypes? (fails with rep\_datatype?)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   540
        * nested datatypes? (fails with datatype\_new?)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   541
    * option 2
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   542
      * \keyw{datatype\_compat}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   543
      * not fully implemented yet?
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   544
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   545
  * register old datatype as new datatype
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   546
    * no clean way yet
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   547
    * if the goal is to do recursion through old datatypes, can register it as
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   548
      a BNF (Section XXX)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   549
    * can also derive destructors etc. using @{command wrap_free_constructors}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   550
      (Section XXX)
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   551
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   552
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   553
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   554
section {* Defining Recursive Functions
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   555
  \label{sec:defining-recursive-functions} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   556
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   557
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   558
This describes how to specify recursive functions over datatypes
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   559
specified using @{command datatype_new}. The focus in on the \keyw{primrec\_new}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   560
command, which supports primitive recursion. A few examples feature the
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   561
@{command fun} and @{command function} commands, described in a separate
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   562
tutorial \cite{isabelle-function}.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   563
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   564
%%% TODO: partial_function?
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   565
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   566
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   567
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   568
More examples in \verb|~~/src/HOL/BNF/Examples|.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   569
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   570
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   571
subsection {* Examples
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   572
  \label{ssec:primrec-examples} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   573
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   574
subsubsection {* Nonrecursive Types *}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   575
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   576
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   577
  * simple (depth 1) pattern matching on the left-hand side
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   578
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   579
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   580
    primrec_new bool_of_trool :: "trool \<Rightarrow> bool" where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   581
      "real_of_trool Faalse = False" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   582
      "real_of_trool Truue = True"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   583
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   584
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   585
  * OK to specify the cases in a different order
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   586
  * OK to leave out some case (but get a warning -- maybe we need a "quiet"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   587
    or "silent" flag?)
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   588
    * case is then unspecified
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   589
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   590
More examples:
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   591
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   592
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   593
    primrec_new list_of_maybe :: "'a maybe \<Rightarrow> 'a list" where
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   594
      "list_of_maybe Nothing = []" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   595
      "list_of_maybe (Just a) = [a]"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   596
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   597
    primrec_new maybe_def :: "'a \<Rightarrow> 'a maybe \<Rightarrow> 'a" where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   598
      "maybe_def d Nothing = d" |
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   599
      "maybe_def _ (Just a) = a"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   600
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   601
    primrec_new mirrror :: "('a, 'b, 'c) triple \<Rightarrow> ('c, 'b, 'a) triple" where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   602
      "mirrror (Triple a b c) = Triple c b a"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   603
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   604
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   605
subsubsection {* Simple Recursion *}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   606
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   607
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   608
again, simple pattern matching on left-hand side, but possibility
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   609
to call a function recursively on an argument to a constructor:
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   610
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   611
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   612
    primrec_new replicate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   613
      "rep 0 _ = []" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   614
      "rep (Suc n) a = a # rep n a"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   615
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   616
text {*
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   617
we don't like the confusing name @{const nth}:
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   618
*}
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   619
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   620
    primrec_new at :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   621
      "at (a # as) j =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   622
         (case j of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   623
            0 \<Rightarrow> a
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   624
          | Suc j' \<Rightarrow> at as j')"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   625
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   626
    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
   627
      "tfold _ (TNil b) = b" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   628
      "tfold f (TCons a as) = f a (tfold f as)"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   629
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   630
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   631
subsubsection {* Mutual Recursion *}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   632
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   633
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   634
E.g., converting even/odd naturals to plain old naturals:
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   635
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   636
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   637
    primrec_new
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   638
      nat_of_enat :: "enat \<Rightarrow> nat" and
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   639
      nat_of_onat :: "onat => nat"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   640
    where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   641
      "nat_of_enat EZero = 0" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   642
      "nat_of_enat (ESuc n) = Suc (nat_of_onat n)" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   643
      "nat_of_onat (OSuc n) = Suc (nat_of_enat n)"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   644
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   645
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   646
Mutual recursion is even possible within a single type, an innovation over the
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   647
old package:
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   648
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   649
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   650
    primrec_new
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   651
      even :: "nat \<Rightarrow> bool" and
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   652
      odd :: "nat \<Rightarrow> bool"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   653
    where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   654
      "even 0 = True" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   655
      "even (Suc n) = odd n" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   656
      "odd 0 = False" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   657
      "odd (Suc n) = even n"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   658
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   659
text {*
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   660
More elaborate:
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   661
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   662
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   663
    primrec_new
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   664
      eval\<^sub>e :: "('a, 'b) exp \<Rightarrow> real" and
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   665
      eval\<^sub>t :: "('a, 'b) trm \<Rightarrow> real" and
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   666
      eval\<^sub>f :: "('a, 'b) fct \<Rightarrow> real"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   667
    where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   668
      "eval\<^sub>e \<gamma> \<xi> (Term t) = eval\<^sub>t \<gamma> \<xi> t" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   669
      "eval\<^sub>e \<gamma> \<xi> (Sum t e) = eval\<^sub>t \<gamma> \<xi> t + eval\<^sub>e \<gamma> \<xi> e" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   670
      "eval\<^sub>t \<gamma> \<xi> (Factor f) = eval\<^sub>f \<gamma> \<xi> f)" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   671
      "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
   672
      "eval\<^sub>f \<gamma> _ (Const a) = \<gamma> a" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   673
      "eval\<^sub>f _ \<xi> (Var b) = \<xi> b" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   674
      "eval\<^sub>f \<gamma> \<xi> (Expr e) = eval\<^sub>e \<gamma> \<xi> e"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   675
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   676
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   677
subsubsection {* Nested Recursion *}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   678
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   679
(*<*)
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   680
    datatype_new 'a treeFF = NodeFF 'a "'a treeFF list"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   681
    datatype_new 'a treeFI = NodeFI 'a "'a treeFI llist"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   682
(*>*)
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   683
    primrec_new atFF0 :: "'a treeFF \<Rightarrow> nat list \<Rightarrow> 'a" where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   684
      "atFF0 (NodeFF a ts) js =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   685
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   686
            [] \<Rightarrow> a
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   687
          | j # js' \<Rightarrow> at (map (\<lambda>t. atFF0 t js') ts) j)"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   688
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   689
    primrec_new atFI :: "'a treeFI \<Rightarrow> nat list \<Rightarrow> 'a" where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   690
      "atFF (NodeFI a ts) js =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   691
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   692
            [] \<Rightarrow> a
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   693
          | j # js' \<Rightarrow> at (llist_map (\<lambda>t. atFF t js') ts) j)"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   694
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   695
    primrec_new sum_btree :: "('a\<Colon>plus) btree \<Rightarrow> 'a" where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   696
      "sum_btree (BNode a lt rt) =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   697
         a + maybe_def 0 (maybe_map sum_btree lt) +
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   698
           maybe_def 0 (maybe_map sum_btree rt)"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   699
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   700
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   701
subsubsection {* Nested-as-Mutual Recursion *}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   702
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   703
text {*
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   704
  * can pretend a nested type is mutually recursive
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   705
  * avoids the higher-order map
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   706
  * e.g.
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   707
*}
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   708
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   709
    primrec_new
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   710
      at_treeFF :: "'a treeFF \<Rightarrow> nat list \<Rightarrow> 'a" and
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   711
      at_treesFF :: "'a treeFF list \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> 'a"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   712
    where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   713
      "at_treeFF (NodeFF a ts) js =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   714
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   715
            [] \<Rightarrow> a
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   716
          | j # js' \<Rightarrow> at_treesFF ts j js')" |
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   717
      "at_treesFF (t # ts) j =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   718
         (case j of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   719
            0 \<Rightarrow> at_treeFF t
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   720
          | Suc j' \<Rightarrow> at_treesFF ts j')"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   721
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   722
    primrec_new
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   723
      sum_btree :: "('a\<Colon>plus) btree \<Rightarrow> 'a" and
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   724
      sum_btree_maybe :: "('a\<Colon>plus) btree maybe \<Rightarrow> 'a"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   725
    where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   726
      "sum_btree (BNode a lt rt) =
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   727
         a + sum_btree_maybe lt + sum_btree_maybe rt" |
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   728
      "sum_btree_maybe Nothing = 0" |
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   729
      "sum_btree_maybe (Just t) = sum_btree t"
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   730
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   731
text {*
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   732
  * this can always be avoided;
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   733
     * e.g. in our previous example, we first mapped the recursive
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   734
       calls, then we used a generic at function to retrieve the result
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   735
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   736
  * there's no hard-and-fast rule of when to use one or the other,
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   737
    just like there's no rule when to use fold and when to use
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   738
    primrec_new
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   739
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   740
  * higher-order approach, considering nesting as nesting, is more
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   741
    compositional -- e.g. we saw how we could reuse an existing polymorphic
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   742
    at or maybe_def, whereas at_treesFF is much more specific
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   743
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   744
  * but:
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   745
     * is perhaps less intuitive, because it requires higher-order thinking
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   746
     * may seem inefficient, and indeed with the code generator the
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   747
       mutually recursive version might be nicer
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   748
     * is somewhat indirect -- must apply a map first, then compute a result
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   749
       (cannot mix)
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   750
     * the auxiliary functions like at_treesFF are sometimes useful in own right
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   751
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   752
  * impact on automation unclear
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   753
*}
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   754
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   755
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   756
subsection {* Syntax
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   757
  \label{ssec:primrec-syntax} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   758
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   759
text {*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   760
Primitive recursive functions have the following general syntax:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   761
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   762
@{rail "
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   763
  @@{command primrec_new} @{syntax target}? @{syntax \"fixes\"} \\ @'where'
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   764
    (@{syntax primrec_equation} + '|')
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   765
  ;
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   766
  @{syntax_def primrec_equation}: @{syntax thmdecl}? @{syntax prop}
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   767
"}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   768
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   769
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   770
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   771
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   772
  \label{ssec:primrec-generated-theorems} *}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   773
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   774
text {*
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   775
  * synthesized nonrecursive definition
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   776
  * user specification is rederived from it, exactly as entered
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   777
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   778
  * induct
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   779
    * mutualized
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   780
    * without some needless induction hypotheses if not used
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   781
  * fold, rec
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   782
    * mutualized
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   783
*}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   784
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   785
subsection {* Recursive Default Values for Selectors
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   786
  \label{ssec:recursive-default-values-for-selectors} *}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   787
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   788
text {*
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   789
A datatype selector @{text un_D} can have a default value for each constructor
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   790
on which it is not otherwise specified. Occasionally, it is useful to have the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   791
default value be defined recursively. This produces a chicken-and-egg situation
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   792
that appears unsolvable, because the datatype is not introduced yet at the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   793
moment when the selectors are introduced. Of course, we can always define the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   794
selectors manually afterward, but we then have to state and prove all the
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   795
characteristic theorems ourselves instead of letting the package do it.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   796
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   797
Fortunately, there is a fairly elegant workaround that relies on overloading and
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   798
that avoids the tedium of manual derivations:
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   799
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   800
\begin{enumerate}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   801
\setlength{\itemsep}{0pt}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   802
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   803
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   804
Introduce a fully unspecified constant @{text "un_D\<^sub>0 \<Colon> 'a"} using
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   805
@{keyword consts}.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   806
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   807
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   808
Define the datatype, specifying @{text "un_D\<^sub>0"} as the selector's default value.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   809
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   810
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   811
Define the behavior of @{text "un_D\<^sub>0"} on values of the newly introduced datatype
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   812
using the @{command overloading} command.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   813
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   814
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   815
Derive the desired equation on @{text un_D} from the characteristic equations
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   816
for @{text "un_D\<^sub>0"}.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   817
\end{enumerate}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   818
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   819
The following example illustrates this procedure:
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   820
*}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   821
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   822
    consts termi\<^sub>0 :: 'a
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   823
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   824
    datatype_new (*<*)(rep_compat) (*>*)('a, 'b) tlist_ =
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   825
      TNil (termi: 'b) (defaults ttl: TNil)
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   826
    | TCons (thd: 'a) (ttl : "('a, 'b) tlist_") (defaults termi: "\<lambda>_ xs. termi\<^sub>0 xs")
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   828
(*<*)
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   829
    (* FIXME: remove hack once "primrec_new" is in place *)
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   830
    rep_datatype TNil TCons
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   831
    by (erule tlist_.induct, assumption) auto
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   832
(*>*)
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   833
    overloading
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   834
      termi\<^sub>0 \<equiv> "termi\<^sub>0 \<Colon> ('a, 'b) tlist_ \<Rightarrow> 'b"
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   835
    begin
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   836
(*<*)
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   837
    (* FIXME: remove hack once "primrec_new" is in place *)
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   838
    fun termi\<^sub>0 :: "('a, 'b) tlist_ \<Rightarrow> 'b" where
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   839
    "termi\<^sub>0 (TNil y) = y" |
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   840
    "termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   841
(*>*)
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   842
    primrec_new termi\<^sub>0 :: "('a, 'b) tlist_ \<Rightarrow> 'b" where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   843
    "termi\<^sub>0 (TNil y) = y" |
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   844
    "termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   845
    end
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   846
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   847
    lemma terminal_TCons[simp]: "termi (TCons x xs) = termi xs"
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   848
    by (cases xs) auto
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   849
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   850
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   851
subsection {* Compatibility Issues
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   852
  \label{ssec:datatype-compatibility-issues} *}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   853
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   854
text {*
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   855
  * different induction in nested case
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   856
    * solution: define nested-as-mutual functions with primrec,
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   857
      and look at .induct
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   858
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   859
  * different induction and recursor in nested case
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   860
    * only matters to low-level users; they can define a dummy function to force
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   861
      generation of mutualized recursor
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   862
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   863
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   864
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   865
section {* Defining Codatatypes
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   866
  \label{sec:defining-codatatypes} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   867
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   868
text {*
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   869
This section describes how to specify codatatypes using the @{command codatatype}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   870
command.
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   871
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   872
  * libraries include some useful codatatypes, notably lazy lists;
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   873
    see the ``Coinductive'' AFP entry \cite{xxx} for an elaborate library
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   874
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   875
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   876
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   877
subsection {* Examples
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   878
  \label{ssec:codatatype-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   879
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   880
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   881
More examples in \verb|~~/src/HOL/BNF/Examples|.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   882
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   883
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   884
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   885
subsection {* Syntax
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   886
  \label{ssec:codatatype-syntax} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   887
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   888
text {*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   889
Definitions of codatatypes have almost exactly the same syntax as for datatypes
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   890
(Section~\ref{ssec:datatype-syntax}), with two exceptions: The command is called
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52868
diff changeset
   891
@{command codatatype}; the \keyw{no\_discs\_sels} option is not available, because
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   892
destructors are a central notion for codatatypes.
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   893
*}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   894
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   895
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   896
  \label{ssec:codatatype-generated-theorems} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   897
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   898
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   899
section {* Defining Corecursive Functions
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   900
  \label{sec:defining-corecursive-functions} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   901
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   902
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   903
This section describes how to specify corecursive functions using the
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   904
\keyw{primcorec} command.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   905
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   906
%%% TODO: partial_function? E.g. for defining tail recursive function on lazy
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   907
%%% lists (cf. terminal0 in TLList.thy)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   908
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   909
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   910
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   911
subsection {* Examples
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   912
  \label{ssec:primcorec-examples} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   913
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   914
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   915
More examples in \verb|~~/src/HOL/BNF/Examples|.
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   916
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   917
Also, for default values, the same trick as for datatypes is possible for
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   918
codatatypes (Section~\ref{ssec:recursive-default-values-for-selectors}).
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   919
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   920
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   921
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   922
subsection {* Syntax
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   923
  \label{ssec:primcorec-syntax} *}
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   924
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   925
text {*
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   926
Primitive corecursive definitions have the following general syntax:
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   927
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   928
@{rail "
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   929
  @@{command primcorec} @{syntax target}? @{syntax \"fixes\"} \\ @'where'
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   930
    (@{syntax primcorec_formula} + '|')
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   931
  ;
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   932
  @{syntax_def primcorec_formula}: @{syntax thmdecl}? @{syntax prop}
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   933
    (@'of' (@{syntax term} * ))?
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   934
"}
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   935
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   936
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   937
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   938
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   939
  \label{ssec:primcorec-generated-theorems} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   940
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   941
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   942
section {* Registering Bounded Natural Functors
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   943
  \label{sec:registering-bounded-natural-functors} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   944
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   945
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   946
This section explains how to set up the (co)datatype package to allow nested
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   947
recursion through custom well-behaved type constructors. The key concept is that
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   948
of a bounded natural functor (BNF).
52829
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
   949
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
   950
    * @{command bnf}
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
   951
    * @{command print_bnfs}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   952
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   953
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   954
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   955
subsection {* Example
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   956
  \label{ssec:bnf-examples} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   957
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   958
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   959
More examples in \verb|~~/src/HOL/BNF/Basic_BNFs.thy| and
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   960
\verb|~~/src/HOL/BNF/More_BNFs.thy|.
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   961
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   962
Mention distinction between live and dead type arguments;
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   963
mention =>.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   964
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   965
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   966
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   967
subsection {* Syntax
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   968
  \label{ssec:bnf-syntax} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   969
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   970
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   971
section {* Generating Free Constructor Theorems
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   972
  \label{sec:generating-free-constructor-theorems} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   973
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   974
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   975
This section explains how to derive convenience theorems for free constructors,
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   976
as performed internally by @{command datatype_new} and @{command codatatype}.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   977
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   978
  * need for this is rare but may arise if you want e.g. to add destructors to
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   979
    a type not introduced by ...
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   980
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   981
  * also useful for compatibility with old package, e.g. add destructors to
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   982
    old @{command datatype}
52829
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
   983
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
   984
  * @{command wrap_free_constructors}
52969
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52868
diff changeset
   985
    * \keyw{no\_discs\_sels}, \keyw{rep\_compat}
f2df0730f8ac clarified option name (since case/fold/rec are also destructors)
blanchet
parents: 52868
diff changeset
   986
    * hack to have both co and nonco view via locale (cf. ext nats)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   987
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   988
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   989
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   990
subsection {* Example
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   991
  \label{ssec:ctors-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   992
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   993
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   994
subsection {* Syntax
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   995
  \label{ssec:ctors-syntax} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   996
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   997
text {*
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   998
Free constructor wrapping has the following general syntax:
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   999
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1000
@{rail "
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1001
  @@{command wrap_free_constructors} @{syntax target}? @{syntax dt_options} \\
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1002
    @{syntax fc_conss} @{syntax name} \\
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1003
    (@{syntax fc_discs} (@{syntax fc_sels} @{syntax fc_sel_defaults}? )? )?
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1004
  ;
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1005
  @{syntax_def fc_conss}: '[' (@{syntax term} + ',') ']'
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1006
  ;
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1007
  @{syntax_def fc_discs}: '[' (@{syntax name} + ',') ']'
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1008
  ;
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1009
  @{syntax_def fc_sels}: '[' ('[' (@{syntax name} + ',') ']' + ',') ']'
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1010
  ;
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1011
  @{syntax_def fc_sel_defaults}: '[' ('[' (@{syntax name} ':' @{syntax term} + ',') ']' + ',') ']'
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1012
"}
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1013
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1014
options: no_discs_sels rep_compat
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  1015
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1016
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1017
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1018
  \label{ssec:ctors-generated-theorems} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1019
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1020
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1021
section {* Standard ML Interface
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1022
  \label{sec:standard-ml-interface} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1023
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1024
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1025
This section describes the package's programmatic interface.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1026
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1027
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1028
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1029
section {* Interoperability
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1030
  \label{sec:interoperability} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1031
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1032
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1033
This section is concerned with the packages' interaction with other Isabelle
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1034
packages and tools, such as the code generator and the counterexample
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1035
generators.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1036
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1037
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1038
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1039
subsection {* Transfer and Lifting
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1040
  \label{ssec:transfer-and-lifting} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1041
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1042
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1043
subsection {* Code Generator
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1044
  \label{ssec:code-generator} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1045
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1046
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1047
subsection {* Quickcheck
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1048
  \label{ssec:quickcheck} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1049
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1050
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1051
subsection {* Nitpick
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1052
  \label{ssec:nitpick} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1053
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1054
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1055
subsection {* Nominal Isabelle
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1056
  \label{ssec:nominal-isabelle} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1057
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1058
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1059
section {* Known Bugs and Limitations
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1060
  \label{sec:known-bugs-and-limitations} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1061
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1062
text {*
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1063
This section lists known open issues of the package.
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1064
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1065
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1066
text {*
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1067
* primrec\_new and primcorec are vaporware
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1068
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1069
* slow n-ary mutual (co)datatype, avoid as much as possible (e.g. using nesting)
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1070
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1071
* issues with HOL-Proofs?
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1072
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1073
* partial documentation
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1074
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
  1075
* too much output by commands like "datatype_new" and "codatatype"
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1076
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1077
* no direct way to define recursive functions for default values -- but show trick
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1078
  based on overloading
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1079
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1080
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1081
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1082
section {* Acknowledgments
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1083
  \label{sec:acknowledgments} *}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1084
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  1085
text {*
52829
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1086
Tobias Nipkow and Makarius Wenzel made this work possible. Andreas Lochbihler
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1087
provided lots of comments on earlier versions of the package, especially for the
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1088
coinductive part. Brian Huffman suggested major simplifications to the internal
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1089
constructions, much of which has yet to be implemented. Florian Haftmann and
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1090
Christian Urban provided general advice advice on Isabelle and package writing.
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1091
Stefan Milius and Lutz Schr\"oder suggested an elegant prove to eliminate one of
591e76f2651e minor doc fixes
blanchet
parents: 52828
diff changeset
  1092
the BNF assumptions.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1093
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1094
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1095
end