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