src/Doc/Datatypes/Datatypes.thy
author desharna
Tue, 24 Jun 2014 13:48:14 +0200
changeset 57304 d2061dc953a4
parent 57206 d9be905d6283
child 57398 882091eb1e9a
permissions -rw-r--r--
document property 'rel_coinduct'
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
57079
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
     3
    Author:     Martin Desharnais, Ecole de technologie superieure
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
     4
    Author:     Lorenz Panny, TU Muenchen
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
     5
    Author:     Andrei Popescu, TU Muenchen
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
     6
    Author:     Dmitriy Traytel, TU Muenchen
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     7
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
     8
Tutorial for (co)datatype definitions with the new package.
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
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    11
theory Datatypes
55073
9b96fb4c8cfd reduced dependencies + updated docs
blanchet
parents: 55029
diff changeset
    12
imports
9b96fb4c8cfd reduced dependencies + updated docs
blanchet
parents: 55029
diff changeset
    13
  Setup
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
    14
  "~~/src/HOL/Library/BNF_Axiomatization"
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
    15
  "~~/src/HOL/Library/Cardinal_Notations"
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
    16
  "~~/src/HOL/Library/FSet"
55073
9b96fb4c8cfd reduced dependencies + updated docs
blanchet
parents: 55029
diff changeset
    17
  "~~/src/HOL/Library/Simps_Case_Conv"
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    18
begin
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    19
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    20
section {* Introduction
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
    21
  \label{sec:introduction} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    22
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    23
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    24
The 2013 edition of Isabelle introduced a new definitional package for freely
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    25
generated datatypes and codatatypes. The datatype support is similar to that
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    26
provided by the earlier package due to Berghofer and Wenzel
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    27
\cite{Berghofer-Wenzel:1999:TPHOL}, documented in the Isar reference manual
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
    28
\cite{isabelle-isar-ref}; indeed, replacing the keyword \keyw{datatype} by
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    29
@{command datatype_new} is usually all that is needed to port existing theories
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    30
to use the new package.
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    31
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    32
Perhaps the main advantage of the new package is that it supports recursion
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
    33
through a large class of non-datatypes, such as finite sets:
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    34
*}
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    35
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
    36
    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
    37
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    38
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    39
\noindent
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
    40
Another strong point is the support for local definitions:
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
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    43
    context linorder
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    44
    begin
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    45
    datatype_new flag = Less | Eq | Greater
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    46
    end
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    47
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    48
text {*
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    49
\noindent
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    50
Furthermore, the package provides a lot of convenience, including automatically
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    51
generated discriminators, selectors, and relators as well as a wealth of
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    52
properties about them.
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    53
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    54
In addition to inductive datatypes, the new package supports coinductive
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    55
datatypes, or \emph{codatatypes}, which allow infinite values. For example, the
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    56
following command introduces the type of lazy lists, which comprises both finite
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    57
and infinite values:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    58
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    59
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    60
(*<*)
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
    61
    locale early
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
    62
    locale late
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
    63
(*>*)
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
    64
    codatatype (*<*)(in early) (*>*)'a llist = LNil | LCons 'a "'a llist"
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    65
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    66
text {*
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    67
\noindent
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    68
Mixed inductive--coinductive recursion is possible via nesting. Compare the
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
    69
following four Rose tree examples:
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    70
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    71
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
    72
    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
    73
    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
    74
    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
    75
    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
    76
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    77
text {*
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    78
The first two tree types allow only paths of finite length, whereas the last two
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    79
allow infinite paths. Orthogonally, the nodes in the first and third types have
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    80
finitely many direct subtrees, whereas those of the second and fourth may have
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
    81
infinite branching.
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
    82
55073
9b96fb4c8cfd reduced dependencies + updated docs
blanchet
parents: 55029
diff changeset
    83
The package is part of @{theory Main}. Additional functionality is provided by
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
    84
the theory @{theory BNF_Axiomatization}, located in the directory
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
    85
\verb|~~/src/HOL/Library|.
55073
9b96fb4c8cfd reduced dependencies + updated docs
blanchet
parents: 55029
diff changeset
    86
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    87
The package, like its predecessor, fully adheres to the LCF philosophy
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    88
\cite{mgordon79}: The characteristic theorems associated with the specified
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    89
(co)datatypes are derived rather than introduced axiomatically.%
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
    90
\footnote{If the @{text quick_and_dirty} option is enabled, some of the
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
    91
internal constructions and most of the internal proof obligations are skipped.}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    92
The package's metatheory is described in a pair of papers
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
    93
\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
    94
\emph{bounded natural functor} (BNF)---a well-behaved type constructor for which
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
    95
nested (co)recursion is supported.
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
    96
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    97
This tutorial is organized as follows:
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
    98
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
    99
\begin{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   100
\setlength{\itemsep}{0pt}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   101
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   102
\item Section \ref{sec:defining-datatypes}, ``Defining Datatypes,''
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   103
describes how to specify datatypes using the @{command datatype_new} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   104
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   105
\item Section \ref{sec:defining-recursive-functions}, ``Defining Recursive
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   106
Functions,'' describes how to specify recursive functions using
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   107
@{command primrec}.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   108
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   109
\item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,''
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   110
describes how to specify codatatypes using the @{command codatatype} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   111
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
   112
\item Section \ref{sec:defining-corecursive-functions}, ``Defining Corecursive
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   113
Functions,'' describes how to specify corecursive functions using the
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
   114
@{command primcorec} and @{command primcorecursive} commands.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   115
55351
blanchet
parents: 55350
diff changeset
   116
\item Section \ref{sec:introducing-bounded-natural-functors}, ``Introducing
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   117
Bounded Natural Functors,'' explains how to use the @{command bnf} command
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   118
to register arbitrary type constructors as BNFs.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   119
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   120
\item Section
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   121
\ref{sec:deriving-destructors-and-theorems-for-free-constructors},
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   122
``Deriving Destructors and Theorems for Free Constructors,'' explains how to
55468
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
   123
use the command @{command free_constructors} to derive destructor constants
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   124
and theorems for freely generated types, as performed internally by @{command
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   125
datatype_new} and @{command codatatype}.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   126
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   127
%\item Section \ref{sec:standard-ml-interface}, ``Standard ML Interface,''
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   128
%describes the package's programmatic interface.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   129
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   130
%\item Section \ref{sec:interoperability}, ``Interoperability,''
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   131
%is concerned with the packages' interaction with other Isabelle packages and
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   132
%tools, such as the code generator and the counterexample generators.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   133
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   134
%\item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   135
%Limitations,'' concludes with known open issues at the time of writing.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   136
\end{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   137
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   138
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   139
\newbox\boxA
54185
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
   140
\setbox\boxA=\hbox{\texttt{NOSPAM}}
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
   141
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
   142
\newcommand\authoremaili{\texttt{blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   143
in.\allowbreak tum.\allowbreak de}}
57079
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   144
\newcommand\authoremailii{\texttt{desh{\color{white}NOSPAM}\kern-\wd\boxA{}arna@\allowbreak
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   145
in.\allowbreak tum.\allowbreak de}}
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   146
\newcommand\authoremailiii{\texttt{lore{\color{white}NOSPAM}\kern-\wd\boxA{}nz.panny@\allowbreak
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   147
in.\allowbreak tum.\allowbreak de}}
57079
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   148
\newcommand\authoremailiv{\texttt{pope{\color{white}NOSPAM}\kern-\wd\boxA{}scua@\allowbreak
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   149
in.\allowbreak tum.\allowbreak de}}
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   150
\newcommand\authoremailv{\texttt{tray{\color{white}NOSPAM}\kern-\wd\boxA{}tel@\allowbreak
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   151
in.\allowbreak tum.\allowbreak de}}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   152
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
   153
The command @{command datatype_new} is expected to replace \keyw{datatype} in a
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
   154
future release. Authors of new theories are encouraged to use the new commands,
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
   155
and maintainers of older theories may want to consider upgrading.
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   156
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   157
Comments and bug reports concerning either the tool or this tutorial should be
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   158
directed to the authors at \authoremaili, \authoremailii, \authoremailiii,
57079
aa7f051ba6ab added fifth member to BNF team
blanchet
parents: 57047
diff changeset
   159
\authoremailiv, and \authoremailv.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   160
*}
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   161
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   162
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   163
section {* Defining Datatypes
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   164
  \label{sec:defining-datatypes} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   165
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   166
text {*
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
   167
Datatypes can be specified using the @{command datatype_new} command.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   168
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
   169
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   170
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   171
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   172
  \label{ssec:datatype-introductory-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   173
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
   174
text {*
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
   175
Datatypes are illustrated through concrete examples featuring different flavors
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
   176
of recursion. More examples can be found in the directory
54185
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
   177
\verb|~~/src/HOL/|\allowbreak\verb|BNF/Examples|.
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
   178
*}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   179
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   180
subsubsection {* Nonrecursive Types
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   181
  \label{sssec:datatype-nonrecursive-types} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   182
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   183
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   184
Datatypes are introduced by specifying the desired names and argument types for
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   185
their constructors. \emph{Enumeration} types are the simplest form of datatype.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   186
All their constructors are nullary:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   187
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   188
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   189
    datatype_new trool = Truue | Faalse | Perhaaps
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   190
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   191
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   192
\noindent
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   193
Here, @{const Truue}, @{const Faalse}, and @{const Perhaaps} have the type @{typ trool}.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   194
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   195
Polymorphic types are possible, such as the following option type, modeled after
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   196
its homologue from the @{theory Option} theory:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   197
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   198
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   199
(*<*)
56995
61855ade6c7e hide more consts to beautify documentation
blanchet
parents: 56992
diff changeset
   200
    hide_const None Some map_option
54958
4933165fd112 do not use wrong constructor in auto-generated proof goal
panny
parents: 54955
diff changeset
   201
    hide_type option
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   202
(*>*)
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   203
    datatype_new 'a option = None | Some 'a
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   204
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   205
text {*
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   206
\noindent
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   207
The constructors are @{text "None :: 'a option"} and
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   208
@{text "Some :: 'a \<Rightarrow> 'a option"}.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   209
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   210
The next example has three type parameters:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   211
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   212
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   213
    datatype_new ('a, 'b, 'c) triple = Triple 'a 'b 'c
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   214
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   215
text {*
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   216
\noindent
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   217
The constructor is
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   218
@{text "Triple :: 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> ('a, 'b, 'c) triple"}.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   219
Unlike in Standard ML, curried constructors are supported. The uncurried variant
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   220
is also possible:
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   221
*}
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   222
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   223
    datatype_new ('a, 'b, 'c) triple\<^sub>u = Triple\<^sub>u "'a * 'b * 'c"
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
   224
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   225
text {*
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   226
\noindent
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   227
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
   228
enclosed in double quotes, as is customary in Isabelle.
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   229
*}
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   230
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   231
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   232
subsubsection {* Simple Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   233
  \label{sssec:datatype-simple-recursion} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   234
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   235
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   236
Natural numbers are the simplest example of a recursive type:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   237
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   238
53332
c97a05a26dd6 Doc improvements
traytel
parents: 53331
diff changeset
   239
    datatype_new nat = Zero | Suc nat
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   240
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   241
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   242
\noindent
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   243
Lists were shown in the introduction. Terminated lists are a variant that
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   244
stores a value of type @{typ 'b} at the very end:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   245
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   246
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   247
    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
   248
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   249
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   250
subsubsection {* Mutual Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   251
  \label{sssec:datatype-mutual-recursion} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   252
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   253
text {*
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   254
\emph{Mutually recursive} types are introduced simultaneously and may refer to
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   255
each other. The example below introduces a pair of types for even and odd
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   256
natural numbers:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   257
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   258
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   259
    datatype_new even_nat = Even_Zero | Even_Suc odd_nat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   260
    and odd_nat = Odd_Suc even_nat
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   261
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   262
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   263
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   264
Arithmetic expressions are defined via terms, terms via factors, and factors via
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   265
expressions:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   266
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   267
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   268
    datatype_new ('a, 'b) exp =
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   269
      Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   270
    and ('a, 'b) trm =
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   271
      Factor "('a, 'b) fct" | Prod "('a, 'b) fct" "('a, 'b) trm"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   272
    and ('a, 'b) fct =
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   273
      Const 'a | Var 'b | Expr "('a, 'b) exp"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   274
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   275
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   276
subsubsection {* Nested Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   277
  \label{sssec:datatype-nested-recursion} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   278
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   279
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   280
\emph{Nested recursion} occurs when recursive occurrences of a type appear under
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   281
a type constructor. The introduction showed some examples of trees with nesting
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
   282
through lists. A more complex example, that reuses our @{type option} type,
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   283
follows:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   284
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   285
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
   286
    datatype_new 'a btree =
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   287
      BNode 'a "'a btree option" "'a btree option"
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   288
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   289
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   290
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   291
Not all nestings are admissible. For example, this command will fail:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   292
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   293
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   294
    datatype_new 'a wrong = W1 | W2 (*<*)'a
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   295
    typ (*>*)"'a wrong \<Rightarrow> 'a"
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   296
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   297
text {*
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   298
\noindent
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   299
The issue is that the function arrow @{text "\<Rightarrow>"} allows recursion
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   300
only through its right-hand side. This issue is inherited by polymorphic
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   301
datatypes defined in terms of~@{text "\<Rightarrow>"}:
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   302
*}
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   303
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
   304
    datatype_new ('a, 'b) fun_copy = Fun "'a \<Rightarrow> 'b"
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   305
    datatype_new 'a also_wrong = W1 | W2 (*<*)'a
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
   306
    typ (*>*)"('a also_wrong, 'a) fun_copy"
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   307
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   308
text {*
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   309
\noindent
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   310
The following definition of @{typ 'a}-branching trees is legal:
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   311
*}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   312
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   313
    datatype_new 'a ftree = FTLeaf 'a | FTNode "'a \<Rightarrow> 'a ftree"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   314
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   315
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   316
\noindent
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   317
And so is the definition of hereditarily finite sets:
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   318
*}
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   319
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   320
    datatype_new hfset = HFSet "hfset fset"
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   321
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   322
text {*
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
   323
\noindent
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   324
In general, type constructors @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   325
allow recursion on a subset of their type arguments @{text 'a\<^sub>1}, \ldots,
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   326
@{text 'a\<^sub>m}. These type arguments are called \emph{live}; the remaining
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   327
type arguments are called \emph{dead}. In @{typ "'a \<Rightarrow> 'b"} and
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
   328
@{typ "('a, 'b) fun_copy"}, the type variable @{typ 'a} is dead and
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
   329
@{typ 'b} is live.
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   330
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   331
Type constructors must be registered as BNFs to have live arguments. This is
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   332
done automatically for datatypes and codatatypes introduced by the @{command
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   333
datatype_new} and @{command codatatype} commands.
55351
blanchet
parents: 55350
diff changeset
   334
Section~\ref{sec:introducing-bounded-natural-functors} explains how to
blanchet
parents: 55350
diff changeset
   335
register arbitrary type constructors as BNFs.
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   336
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   337
Here is another example that fails:
52806
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   338
*}
146ce45c3619 more work on (co)datatype docs
blanchet
parents: 52805
diff changeset
   339
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   340
    datatype_new 'a pow_list = PNil 'a (*<*)'a
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   341
    datatype_new 'a pow_list' = PNil' 'a (*>*)| PCons "('a * 'a) pow_list"
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   342
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   343
text {*
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   344
\noindent
55351
blanchet
parents: 55350
diff changeset
   345
This attempted definition features a different flavor of nesting, where the
blanchet
parents: 55350
diff changeset
   346
recursive call in the type specification occurs around (rather than inside)
blanchet
parents: 55350
diff changeset
   347
another type constructor.
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   348
*}
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   349
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   350
subsubsection {* Auxiliary Constants and Properties
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   351
  \label{sssec:datatype-auxiliary-constants-and-properties} *}
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
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   357
custom names. In the example below, the familiar names @{text null}, @{text hd},
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   358
@{text tl}, @{text set}, @{text map}, and @{text list_all2}, override the
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   359
default names @{text is_Nil}, @{text un_Cons1}, @{text un_Cons2},
54491
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
   360
@{text set_list}, @{text map_list}, and @{text rel_list}:
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
56995
61855ade6c7e hide more consts to beautify documentation
blanchet
parents: 56992
diff changeset
   373
    hide_const Nil Cons hd tl set map list_all2 rec_list size_list
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
(*>*)
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   377
    datatype_new (set: 'a) list =
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   378
      null: Nil
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   379
    | Cons (hd: 'a) (tl: "'a list")
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   380
    for
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   381
      map: map
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   382
      rel: list_all2
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   383
    where
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   384
      "tl Nil = Nil"
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   385
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   386
text {*
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   387
\noindent
55353
blanchet
parents: 55351
diff changeset
   388
The types of the constants that appear in the specification are listed below.
55351
blanchet
parents: 55350
diff changeset
   389
blanchet
parents: 55350
diff changeset
   390
\medskip
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   391
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   392
\begin{tabular}{@ {}ll@ {}}
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   393
Constructors: &
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   394
  @{text "Nil \<Colon> 'a list"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   395
&
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   396
  @{text "Cons \<Colon> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   397
Discriminator: &
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   398
  @{text "null \<Colon> 'a list \<Rightarrow> bool"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   399
Selectors: &
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   400
  @{text "hd \<Colon> 'a list \<Rightarrow> 'a"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   401
&
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   402
  @{text "tl \<Colon> 'a list \<Rightarrow> 'a list"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   403
Set function: &
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   404
  @{text "set \<Colon> 'a list \<Rightarrow> 'a set"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   405
Map function: &
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   406
  @{text "map \<Colon> ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \\
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   407
Relator: &
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   408
  @{text "list_all2 \<Colon> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool"}
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   409
\end{tabular}
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   410
55351
blanchet
parents: 55350
diff changeset
   411
\medskip
blanchet
parents: 55350
diff changeset
   412
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   413
The discriminator @{const null} and the selectors @{const hd} and @{const tl}
55351
blanchet
parents: 55350
diff changeset
   414
are characterized by the following conditional equations:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   415
%
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   416
\[@{thm list.collapse(1)[of xs, no_vars]}
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   417
  \qquad @{thm list.collapse(2)[of xs, no_vars]}\]
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   418
%
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   419
For two-constructor datatypes, a single discriminator constant is sufficient.
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
   420
The discriminator associated with @{const Cons} is simply
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
   421
@{term "\<lambda>xs. \<not> null xs"}.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   422
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   423
The \keyw{where} clause at the end of the command specifies a default value for
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   424
selectors applied to constructors on which they are not a priori specified.
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   425
Here, it is used to ensure that the tail of the empty list is itself (instead of
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   426
being left unspecified).
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   427
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   428
Because @{const Nil} is nullary, it is also possible to use
57091
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   429
@{term "\<lambda>xs. xs = Nil"} as a discriminator. This is the default behavior
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   430
if we omit the identifier @{const null} and the associated colon. Some users
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   431
argue against this, because the mixture of constructors and selectors in the
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   432
characteristic theorems can lead Isabelle's automation to switch between the
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   433
constructor and the destructor view in surprising ways.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   434
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   435
The usual mixfix syntax annotations are available for both types and
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   436
constructors. For example:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
   437
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   438
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   439
(*<*)
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   440
    end
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
   441
(*>*)
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   442
    datatype_new ('a, 'b) prod (infixr "*" 20) = Pair 'a 'b
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   443
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   444
text {* \blankline *}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   445
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   446
    datatype_new (set: 'a) list =
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   447
      null: Nil ("[]")
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   448
    | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65)
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   449
    for
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   450
      map: map
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   451
      rel: list_all2
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   452
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   453
text {*
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   454
\noindent
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   455
Incidentally, this is how the traditional syntax can be set up:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   456
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   457
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   458
    syntax "_list" :: "args \<Rightarrow> 'a list" ("[(_)]")
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   459
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   460
text {* \blankline *}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   461
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   462
    translations
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   463
      "[x, xs]" == "x # [xs]"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
   464
      "[x]" == "x # []"
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   465
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   466
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   467
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   468
  \label{ssec:datatype-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   469
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   470
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   471
subsubsection {* \keyw{datatype\_new}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   472
  \label{sssec:datatype-new} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   473
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   474
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   475
\begin{matharray}{rcl}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   476
  @{command_def "datatype_new"} & : & @{text "local_theory \<rightarrow> local_theory"}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   477
\end{matharray}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   478
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   479
@{rail \<open>
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 54958
diff changeset
   480
  @@{command datatype_new} target? @{syntax dt_options}? \<newline>
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   481
    (@{syntax dt_name} '=' (@{syntax dt_ctor} + '|') \<newline>
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   482
     @{syntax map_rel}? (@'where' (prop + '|'))? + @'and')
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   483
  ;
57094
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
   484
  @{syntax_def dt_options}: '(' (('discs_sels' | 'no_code') + ',') ')'
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   485
  ;
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   486
  @{syntax_def map_rel}: @'for' ((('map' | 'rel') ':' name) +)
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   487
\<close>}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   488
55351
blanchet
parents: 55350
diff changeset
   489
\medskip
blanchet
parents: 55350
diff changeset
   490
blanchet
parents: 55350
diff changeset
   491
\noindent
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
   492
The @{command datatype_new} command introduces a set of mutually recursive
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
   493
datatypes specified by their constructors.
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
   494
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   495
The syntactic entity \synt{target} can be used to specify a local
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   496
context (e.g., @{text "(in linorder)"} \cite{isabelle-isar-ref}),
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   497
and \synt{prop} denotes a HOL proposition.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
   498
56124
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
   499
The optional target is optionally followed by one or both of the following
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
   500
options:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   501
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   502
\begin{itemize}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   503
\setlength{\itemsep}{0pt}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   504
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   505
\item
57094
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
   506
The @{text "discs_sels"} option indicates that discriminators and selectors
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
   507
should be generated. The option is implicitly enabled if names are specified for
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
   508
discriminators or selectors.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   509
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   510
\item
54626
8a5e82425e55 added 'no_code' option
blanchet
parents: 54624
diff changeset
   511
The @{text "no_code"} option indicates that the datatype should not be
8a5e82425e55 added 'no_code' option
blanchet
parents: 54624
diff changeset
   512
registered for code generation.
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   513
\end{itemize}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   514
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   515
The optional \keyw{where} clause specifies default values for selectors.
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   516
Each proposition must be an equation of the form
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   517
@{text "un_D (C \<dots>) = \<dots>"}, where @{text C} is a constructor and
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   518
@{text un_D} is a selector.
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   519
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   520
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
   521
define, its type parameters, and additional information:
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   522
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   523
@{rail \<open>
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   524
  @{syntax_def dt_name}: @{syntax tyargs}? name mixfix?
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   525
  ;
57092
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
   526
  @{syntax_def tyargs}: typefree | '(' (('dead' | name ':')? typefree + ',') ')'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   527
\<close>}
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   528
55351
blanchet
parents: 55350
diff changeset
   529
\medskip
blanchet
parents: 55350
diff changeset
   530
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   531
\noindent
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   532
The syntactic entity \synt{name} denotes an identifier, \synt{mixfix} denotes
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   533
the usual parenthesized mixfix notation, and \synt{typefree} denotes fixed type
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   534
variable (@{typ 'a}, @{typ 'b}, \ldots) \cite{isabelle-isar-ref}.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   535
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   536
The optional names preceding the type variables allow to override the default
55705
a98a045a6169 updated docs
blanchet
parents: 55641
diff changeset
   537
names of the set functions (@{text set1_t}, \ldots, @{text setM_t}). Type
57092
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
   538
arguments can be marked as dead by entering ``@{text dead}'' in front of the
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
   539
type variable (e.g., ``@{text "(dead 'a)"}''); otherwise, they are live or dead
55705
a98a045a6169 updated docs
blanchet
parents: 55641
diff changeset
   540
(and a set function is generated or not) depending on where they occur in the
a98a045a6169 updated docs
blanchet
parents: 55641
diff changeset
   541
right-hand sides of the definition. Declaring a type argument as dead can speed
a98a045a6169 updated docs
blanchet
parents: 55641
diff changeset
   542
up the type definition but will prevent any later (co)recursion through that
a98a045a6169 updated docs
blanchet
parents: 55641
diff changeset
   543
type argument.
a98a045a6169 updated docs
blanchet
parents: 55641
diff changeset
   544
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   545
Inside a mutually recursive specification, all defined datatypes must
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   546
mention exactly the same type variables in the same order.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   547
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   548
@{rail \<open>
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
   549
  @{syntax_def dt_ctor}: (name ':')? name (@{syntax dt_ctor_arg} * ) mixfix?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   550
\<close>}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   551
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   552
\medskip
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   553
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   554
\noindent
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
   555
The main constituents of a constructor specification are the name of the
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   556
constructor and the list of its argument types. An optional discriminator name
57091
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   557
can be supplied at the front to override the default, which is
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   558
@{text "\<lambda>x. x = C\<^sub>j"} for nullary constructors and
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
   559
@{text t.is_C\<^sub>j} otherwise.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   560
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   561
@{rail \<open>
55472
990651bfc65b updated docs to reflect the new 'free_constructors' syntax
blanchet
parents: 55468
diff changeset
   562
  @{syntax_def dt_ctor_arg}: type | '(' name ':' type ')'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   563
\<close>}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   564
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   565
\medskip
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   566
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   567
\noindent
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   568
The syntactic entity \synt{type} denotes a HOL type \cite{isabelle-isar-ref}.
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   569
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   570
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
   571
name for the corresponding selector to override the default name
53554
78fe0002024d more (co)data docs
blanchet
parents: 53553
diff changeset
   572
(@{text un_C\<^sub>ji}). The same selector names can be reused for several
78fe0002024d more (co)data docs
blanchet
parents: 53553
diff changeset
   573
constructors as long as they share the same type.
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   574
*}
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
   575
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   576
56644
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
   577
subsubsection {* \keyw{datatype\_compat}
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
   578
  \label{sssec:datatype-compat} *}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   579
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   580
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   581
\begin{matharray}{rcl}
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
   582
  @{command_def "datatype_compat"} & : & @{text "local_theory \<rightarrow> local_theory"}
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   583
\end{matharray}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   584
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   585
@{rail \<open>
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
   586
  @@{command datatype_compat} (name +)
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
   587
\<close>}
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   588
55351
blanchet
parents: 55350
diff changeset
   589
\medskip
blanchet
parents: 55350
diff changeset
   590
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
   591
\noindent
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
   592
The @{command datatype_compat} command registers new-style datatypes as
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   593
old-style datatypes. For example:
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   594
*}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   595
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
   596
    datatype_compat even_nat odd_nat
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   597
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   598
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   599
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   600
    ML {* Datatype_Data.get_info @{theory} @{type_name even_nat} *}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   601
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   602
text {*
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   603
The syntactic entity \synt{name} denotes an identifier \cite{isabelle-isar-ref}.
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   604
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   605
The command is interesting because the old datatype package provides some
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   606
functionality that is not yet replicated in the new package, such as the
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   607
integration with Quickcheck.
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   608
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   609
A few remarks concern nested recursive datatypes:
53748
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   610
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   611
\begin{itemize}
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   612
\setlength{\itemsep}{0pt}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   613
55871
a28817253b31 removed (co)iterators from documentation
blanchet
parents: 55705
diff changeset
   614
\item The old-style, nested-as-mutual induction rule and recursor theorems are
a28817253b31 removed (co)iterators from documentation
blanchet
parents: 55705
diff changeset
   615
generated under their usual names but with ``@{text "compat_"}'' prefixed
a28817253b31 removed (co)iterators from documentation
blanchet
parents: 55705
diff changeset
   616
(e.g., @{text compat_tree.induct}).
53748
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   617
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   618
\item All types through which recursion takes place must be new-style datatypes
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   619
or the function type. In principle, it should be possible to support old-style
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   620
datatypes as well, but this has not been implemented yet (and there is currently
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
   621
no way to register old-style datatypes as new-style datatypes).
54184
3fe73f3c84a2 more docs
blanchet
parents: 54183
diff changeset
   622
3fe73f3c84a2 more docs
blanchet
parents: 54183
diff changeset
   623
\item The recursor produced for types that recurse through functions has a
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   624
different signature than with the old package. This might affect the behavior of
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   625
some third-party extensions.
53748
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   626
\end{itemize}
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   627
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
   628
An alternative to @{command datatype_compat} is to use the old package's
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
   629
\keyw{rep\_\allowbreak datatype} command. The associated proof obligations must then be
53748
be0ddf3dd01b updated docs
blanchet
parents: 53722
diff changeset
   630
discharged manually.
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   631
*}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   632
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   633
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   634
subsection {* Generated Constants
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   635
  \label{ssec:datatype-generated-constants} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   636
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   637
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   638
Given a datatype @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   639
with $m > 0$ live type variables and $n$ constructors
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   640
@{text "t.C\<^sub>1"}, \ldots, @{text "t.C\<^sub>n"}, the
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   641
following auxiliary constants are introduced:
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   642
55353
blanchet
parents: 55351
diff changeset
   643
\medskip
blanchet
parents: 55351
diff changeset
   644
blanchet
parents: 55351
diff changeset
   645
\begin{tabular}{@ {}ll@ {}}
blanchet
parents: 55351
diff changeset
   646
Case combinator: &
blanchet
parents: 55351
diff changeset
   647
  @{text t.case_t} (rendered using the familiar @{text case}--@{text of} syntax) \\
blanchet
parents: 55351
diff changeset
   648
Discriminators: &
blanchet
parents: 55351
diff changeset
   649
  @{text "t.is_C\<^sub>1"}$, \ldots, $@{text "t.is_C\<^sub>n"} \\
blanchet
parents: 55351
diff changeset
   650
Selectors: &
blanchet
parents: 55351
diff changeset
   651
  @{text t.un_C\<^sub>11}$, \ldots, $@{text t.un_C\<^sub>1k\<^sub>1} \\
blanchet
parents: 55351
diff changeset
   652
& \quad\vdots \\
blanchet
parents: 55351
diff changeset
   653
& @{text t.un_C\<^sub>n1}$, \ldots, $@{text t.un_C\<^sub>nk\<^sub>n} \\
blanchet
parents: 55351
diff changeset
   654
Set functions: &
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   655
  @{text t.set1_t}, \ldots, @{text t.setm_t} \\
55353
blanchet
parents: 55351
diff changeset
   656
Map function: &
blanchet
parents: 55351
diff changeset
   657
  @{text t.map_t} \\
blanchet
parents: 55351
diff changeset
   658
Relator: &
blanchet
parents: 55351
diff changeset
   659
  @{text t.rel_t} \\
blanchet
parents: 55351
diff changeset
   660
Recursor: &
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   661
  @{text t.rec_t} \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   662
Size function: &
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   663
  @{text t.size_t}
55353
blanchet
parents: 55351
diff changeset
   664
\end{tabular}
blanchet
parents: 55351
diff changeset
   665
blanchet
parents: 55351
diff changeset
   666
\medskip
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   667
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   668
\noindent
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   669
The case combinator, discriminators, and selectors are collectively called
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   670
\emph{destructors}. The prefix ``@{text "t."}'' is an optional component of the
54491
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
   671
names and is normally hidden.
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   672
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   673
In addition to the above, the @{class size} class is instantiated to overload the
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   674
@{const size} function based on @{text t.size_t}.
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   675
*}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   676
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
   677
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   678
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
   679
  \label{ssec:datatype-generated-theorems} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   680
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   681
text {*
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   682
The characteristic theorems generated by @{command datatype_new} are grouped in
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   683
three broad categories:
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   684
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   685
\begin{itemize}
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   686
\setlength{\itemsep}{0pt}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   687
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   688
\item The \emph{free constructor theorems}
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   689
(Section~\ref{sssec:free-constructor-theorems}) are properties of the
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   690
constructors and destructors that can be derived for any freely generated type.
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   691
Internally, the derivation is performed by @{command free_constructors}.
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   692
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   693
\item The \emph{functorial theorems} (Section~\ref{sssec:functorial-theorems})
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   694
are properties of datatypes related to their BNF nature.
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   695
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   696
\item The \emph{inductive theorems} (Section~\ref{sssec:inductive-theorems})
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   697
are properties of datatypes related to their inductive nature.
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   698
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   699
\end{itemize}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   700
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   701
\noindent
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   702
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
   703
command \keyw{print\_theorems} immediately after the datatype definition.
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   704
This list normally excludes low-level theorems that reveal internal
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   705
constructions. To make these accessible, add the line
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   706
*}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   707
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   708
    declare [[bnf_note_all]]
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   709
(*<*)
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   710
    declare [[bnf_note_all = false]]
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
   711
(*>*)
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   712
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   713
text {*
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   714
\noindent
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   715
to the top of the theory file.
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   716
*}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   717
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   718
subsubsection {* Free Constructor Theorems
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   719
  \label{sssec:free-constructor-theorems} *}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   720
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   721
(*<*)
53837
65c775430caa updated docs
blanchet
parents: 53831
diff changeset
   722
    consts nonnull :: 'a
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   723
(*>*)
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   724
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   725
text {*
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   726
The free constructor theorems are partitioned in three subgroups. The first
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   727
subgroup of properties is concerned with the constructors. They are listed below
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   728
for @{typ "'a list"}:
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   729
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   730
\begin{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   731
\begin{description}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   732
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   733
\item[@{text "t."}\hthm{inject} @{text "[iff, induct_simp]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   734
@{thm list.inject[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   735
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   736
\item[@{text "t."}\hthm{distinct} @{text "[simp, induct_simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   737
@{thm list.distinct(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   738
@{thm list.distinct(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   739
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   740
\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
   741
@{thm list.exhaust[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   742
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   743
\item[@{text "t."}\hthm{nchotomy}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   744
@{thm list.nchotomy[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   745
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   746
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   747
\end{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   748
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   749
\noindent
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   750
In addition, these nameless theorems are registered as safe elimination rules:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   751
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   752
\begin{indentblock}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   753
\begin{description}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   754
54386
3514fdfdf59b minor doc fix
blanchet
parents: 54287
diff changeset
   755
\item[@{text "t."}\hthm{distinct {\upshape[}THEN notE}@{text ", elim!"}\hthm{\upshape]}\rm:] ~ \\
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   756
@{thm list.distinct(1)[THEN notE, elim!, no_vars]} \\
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   757
@{thm list.distinct(2)[THEN notE, elim!, no_vars]}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   758
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   759
\end{description}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   760
\end{indentblock}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   761
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   762
\noindent
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   763
The next subgroup is concerned with the case combinator:
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   764
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   765
\begin{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   766
\begin{description}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   767
53798
6a4e3299dfd1 set [code] on case equations
blanchet
parents: 53753
diff changeset
   768
\item[@{text "t."}\hthm{case} @{text "[simp, code]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   769
@{thm list.case(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   770
@{thm list.case(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   771
55398
67e9fdd9ae9e updated docs
blanchet
parents: 55355
diff changeset
   772
\item[@{text "t."}\hthm{case\_cong} @{text "[fundef_cong]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   773
@{thm list.case_cong[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   774
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   775
\item[@{text "t."}\hthm{weak\_case\_cong} @{text "[cong]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   776
@{thm list.weak_case_cong[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   777
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   778
\item[@{text "t."}\hthm{split}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   779
@{thm list.split[no_vars]}
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{split\_asm}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   782
@{thm list.split_asm[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   783
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   784
\item[@{text "t."}\hthm{splits} = @{text "split split_asm"}]
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   785
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   786
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   787
\end{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   788
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   789
\noindent
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   790
The third subgroup revolves around discriminators and selectors:
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   791
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   792
\begin{indentblock}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   793
\begin{description}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   794
53694
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   795
\item[@{text "t."}\hthm{disc} @{text "[simp]"}\rm:] ~ \\
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   796
@{thm list.disc(1)[no_vars]} \\
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   797
@{thm list.disc(2)[no_vars]}
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   798
53703
0c565fec9c78 updated docs
blanchet
parents: 53694
diff changeset
   799
\item[@{text "t."}\hthm{discI}\rm:] ~ \\
0c565fec9c78 updated docs
blanchet
parents: 53694
diff changeset
   800
@{thm list.discI(1)[no_vars]} \\
0c565fec9c78 updated docs
blanchet
parents: 53694
diff changeset
   801
@{thm list.discI(2)[no_vars]}
0c565fec9c78 updated docs
blanchet
parents: 53694
diff changeset
   802
53805
4163160853fd added [code] to selectors
blanchet
parents: 53798
diff changeset
   803
\item[@{text "t."}\hthm{sel} @{text "[simp, code]"}\rm:] ~ \\
53694
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   804
@{thm list.sel(1)[no_vars]} \\
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   805
@{thm list.sel(2)[no_vars]}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   806
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   807
\item[@{text "t."}\hthm{collapse} @{text "[simp]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   808
@{thm list.collapse(1)[no_vars]} \\
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   809
@{thm list.collapse(2)[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   810
53837
65c775430caa updated docs
blanchet
parents: 53831
diff changeset
   811
\item[@{text "t."}\hthm{disc\_exclude} @{text "[dest]"}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   812
These properties are missing for @{typ "'a list"} because there is only one
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   813
proper discriminator. Had the datatype been introduced with a second
53837
65c775430caa updated docs
blanchet
parents: 53831
diff changeset
   814
discriminator called @{const nonnull}, they would have read thusly: \\[\jot]
65c775430caa updated docs
blanchet
parents: 53831
diff changeset
   815
@{prop "null list \<Longrightarrow> \<not> nonnull list"} \\
65c775430caa updated docs
blanchet
parents: 53831
diff changeset
   816
@{prop "nonnull list \<Longrightarrow> \<not> null list"}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   817
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   818
\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
   819
@{thm list.disc_exhaust[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   820
53916
37c31a619eee generate "sel_exhaust" theorem
blanchet
parents: 53863
diff changeset
   821
\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
   822
@{thm list.sel_exhaust[no_vars]}
37c31a619eee generate "sel_exhaust" theorem
blanchet
parents: 53863
diff changeset
   823
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   824
\item[@{text "t."}\hthm{expand}\rm:] ~ \\
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   825
@{thm list.expand[no_vars]}
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   826
53917
bf74357f91f8 generate "sel_splits(_asm)" theorems
blanchet
parents: 53916
diff changeset
   827
\item[@{text "t."}\hthm{sel\_split}\rm:] ~ \\
bf74357f91f8 generate "sel_splits(_asm)" theorems
blanchet
parents: 53916
diff changeset
   828
@{thm list.sel_split[no_vars]}
bf74357f91f8 generate "sel_splits(_asm)" theorems
blanchet
parents: 53916
diff changeset
   829
bf74357f91f8 generate "sel_splits(_asm)" theorems
blanchet
parents: 53916
diff changeset
   830
\item[@{text "t."}\hthm{sel\_split\_asm}\rm:] ~ \\
bf74357f91f8 generate "sel_splits(_asm)" theorems
blanchet
parents: 53916
diff changeset
   831
@{thm list.sel_split_asm[no_vars]}
bf74357f91f8 generate "sel_splits(_asm)" theorems
blanchet
parents: 53916
diff changeset
   832
54491
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
   833
\item[@{text "t."}\hthm{case\_eq\_if}\rm:] ~ \\
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
   834
@{thm list.case_eq_if[no_vars]}
53543
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   835
c6c8dce7e9ab more (co)data docs
blanchet
parents: 53542
diff changeset
   836
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   837
\end{indentblock}
54152
f15bd1f81220 updated docs
blanchet
parents: 54146
diff changeset
   838
f15bd1f81220 updated docs
blanchet
parents: 54146
diff changeset
   839
\noindent
f15bd1f81220 updated docs
blanchet
parents: 54146
diff changeset
   840
In addition, equational versions of @{text t.disc} are registered with the @{text "[code]"}
f15bd1f81220 updated docs
blanchet
parents: 54146
diff changeset
   841
attribute.
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   842
*}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   843
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   844
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   845
subsubsection {* Functorial Theorems
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   846
  \label{sssec:functorial-theorems} *}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   847
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   848
text {*
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   849
The functorial theorems are partitioned in two subgroups. The first subgroup
56992
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   850
consists of properties involving the constructors or the destructors and either
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   851
a set function, the map function, or the relator:
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   852
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   853
\begin{indentblock}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   854
\begin{description}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   855
53798
6a4e3299dfd1 set [code] on case equations
blanchet
parents: 53753
diff changeset
   856
\item[@{text "t."}\hthm{set} @{text "[simp, code]"}\rm:] ~ \\
53694
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   857
@{thm list.set(1)[no_vars]} \\
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   858
@{thm list.set(2)[no_vars]}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   859
56992
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   860
\item[@{text "t."}\hthm{set\_empty}\rm:] ~ \\
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   861
@{thm list.set_empty[no_vars]}
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   862
57153
690cf0d83162 document property 'sel_set'
desharna
parents: 57094
diff changeset
   863
\item[@{text "t."}\hthm{sel\_set}\rm:] ~ \\
690cf0d83162 document property 'sel_set'
desharna
parents: 57094
diff changeset
   864
@{thm list.sel_set[no_vars]}
690cf0d83162 document property 'sel_set'
desharna
parents: 57094
diff changeset
   865
53798
6a4e3299dfd1 set [code] on case equations
blanchet
parents: 53753
diff changeset
   866
\item[@{text "t."}\hthm{map} @{text "[simp, code]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   867
@{thm list.map(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   868
@{thm list.map(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   869
56992
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   870
\item[@{text "t."}\hthm{disc\_map\_iff} @{text "[simp]"}\rm:] ~ \\
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   871
@{thm list.disc_map_iff[no_vars]}
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   872
57047
90d4db566e12 document property 'sel_map'
desharna
parents: 56995
diff changeset
   873
\item[@{text "t."}\hthm{sel\_map}\rm:] ~ \\
90d4db566e12 document property 'sel_map'
desharna
parents: 56995
diff changeset
   874
@{thm list.sel_map[no_vars]}
90d4db566e12 document property 'sel_map'
desharna
parents: 56995
diff changeset
   875
54146
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
   876
\item[@{text "t."}\hthm{rel\_inject} @{text "[simp]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   877
@{thm list.rel_inject(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   878
@{thm list.rel_inject(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   879
54146
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
   880
\item[@{text "t."}\hthm{rel\_distinct} @{text "[simp]"}\rm:] ~ \\
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   881
@{thm list.rel_distinct(1)[no_vars]} \\
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   882
@{thm list.rel_distinct(2)[no_vars]}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   883
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   884
\end{description}
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   885
\end{indentblock}
54146
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
   886
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
   887
\noindent
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
   888
In addition, equational versions of @{text t.rel_inject} and @{text
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
   889
rel_distinct} are registered with the @{text "[code]"} attribute.
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   890
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   891
The second subgroup consists of more abstract properties of the set functions,
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   892
the map function, and the relator:
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   893
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   894
\begin{indentblock}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   895
\begin{description}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   896
56992
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   897
\item[@{text "t."}\hthm{set\_map}\rm:] ~ \\
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   898
@{thm list.set_map[no_vars]}
d0e5225601d3 document property 'disc_map_iff'
desharna
parents: 56957
diff changeset
   899
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   900
\item[@{text "t."}\hthm{map\_comp}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   901
@{thm list.map_cong0[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   902
54624
36301c99ed26 revert making 'map_cong' a 'cong' -- it breaks too many proofs in the AFP
blanchet
parents: 54621
diff changeset
   903
\item[@{text "t."}\hthm{map\_cong} @{text "[fundef_cong]"}\rm:] ~ \\
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   904
@{thm list.map_cong[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   905
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   906
\item[@{text "t."}\hthm{map\_id}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   907
@{thm list.map_id[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   908
56955
ddcfa5d19c1a document 'map_id0'
desharna
parents: 56948
diff changeset
   909
\item[@{text "t."}\hthm{map\_id0}\rm:] ~ \\
ddcfa5d19c1a document 'map_id0'
desharna
parents: 56948
diff changeset
   910
@{thm list.map_id0[no_vars]}
ddcfa5d19c1a document 'map_id0'
desharna
parents: 56948
diff changeset
   911
56904
823f1c979580 Documented new property
desharna
parents: 56655
diff changeset
   912
\item[@{text "t."}\hthm{map\_ident}\rm:] ~ \\
823f1c979580 Documented new property
desharna
parents: 56655
diff changeset
   913
@{thm list.map_ident[no_vars]}
823f1c979580 Documented new property
desharna
parents: 56655
diff changeset
   914
54621
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   915
\item[@{text "t."}\hthm{rel\_compp}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   916
@{thm list.rel_compp[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   917
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   918
\item[@{text "t."}\hthm{rel\_conversep}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   919
@{thm list.rel_conversep[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   920
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   921
\item[@{text "t."}\hthm{rel\_eq}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   922
@{thm list.rel_eq[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   923
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   924
\item[@{text "t."}\hthm{rel\_flip}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   925
@{thm list.rel_flip[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   926
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   927
\item[@{text "t."}\hthm{rel\_mono}\rm:] ~ \\
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   928
@{thm list.rel_mono[no_vars]}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   929
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   930
\end{description}
82a78bc9fa0d docs for forgotten BNF theorems
blanchet
parents: 54616
diff changeset
   931
\end{indentblock}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   932
*}
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   933
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   934
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   935
subsubsection {* Inductive Theorems
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
   936
  \label{sssec:inductive-theorems} *}
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   937
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
   938
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
   939
The inductive theorems are as follows:
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   940
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   941
\begin{indentblock}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   942
\begin{description}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   943
57304
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
   944
\item[@{text "t."}\hthm{induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n, induct t]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   945
@{thm list.induct[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   946
53642
05ca82603671 more (co)data docs
blanchet
parents: 53623
diff changeset
   947
\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
   948
Given $m > 1$ mutually recursive datatypes, this induction rule can be used to
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   949
prove $m$ properties simultaneously.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   950
53798
6a4e3299dfd1 set [code] on case equations
blanchet
parents: 53753
diff changeset
   951
\item[@{text "t."}\hthm{rec} @{text "[simp, code]"}\rm:] ~ \\
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   952
@{thm list.rec(1)[no_vars]} \\
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   953
@{thm list.rec(2)[no_vars]}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   954
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   955
\item[@{text "t."}\hthm{rec\_o\_map}\rm:] ~ \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   956
@{thm list.rec_o_map[no_vars]}
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   957
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   958
\item[@{text "t."}\hthm{size} @{text "[simp, code]"}\rm:] ~ \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   959
@{thm list.size(1)[no_vars]} \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   960
@{thm list.size(2)[no_vars]} \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   961
@{thm list.size(3)[no_vars]} \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   962
@{thm list.size(4)[no_vars]}
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   963
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   964
\item[@{text "t."}\hthm{size\_o\_map}\rm:] ~ \\
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   965
@{thm list.size_o_map[no_vars]}
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   966
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   967
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   968
\end{indentblock}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   969
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   970
\noindent
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   971
For convenience, @{command datatype_new} also provides the following collection:
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   972
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   973
\begin{indentblock}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   974
\begin{description}
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   975
55871
a28817253b31 removed (co)iterators from documentation
blanchet
parents: 55705
diff changeset
   976
\item[@{text "t."}\hthm{simps} = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.rec} @{text t.map} @{text t.rel_inject}] ~ \\
53694
7b453b619b5f use singular to avoid confusion
blanchet
parents: 53691
diff changeset
   977
@{text t.rel_distinct} @{text t.set}
53544
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   978
2176a7e40786 more (co)data docs
blanchet
parents: 53543
diff changeset
   979
\end{description}
53552
eed6efba4e3f more (co)datatype docs
blanchet
parents: 53544
diff changeset
   980
\end{indentblock}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   981
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   982
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   983
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
   984
subsection {* Compatibility Issues
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
   985
  \label{ssec:datatype-compatibility-issues} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
   986
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
   987
text {*
53997
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
   988
The command @{command datatype_new} has been designed to be highly compatible
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
   989
with the old \keyw{datatype}, to ease migration. There are nonetheless a few
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   990
incompatibilities that may arise when porting to the new package:
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   991
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   992
\begin{itemize}
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   993
\setlength{\itemsep}{0pt}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
   994
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   995
\item \emph{The Standard ML interfaces are different.} Tools and extensions
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
   996
written to call the old ML interfaces will need to be adapted to the new
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   997
interfaces. This concerns Quickcheck in particular. Whenever possible, it is
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   998
recommended to use @{command datatype_compat} to register new-style datatypes
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
   999
as old-style datatypes.
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1000
56644
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1001
\item \emph{The constants @{text t_case}, @{text t_rec}, and @{text t_size} are
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1002
now called @{text case_t}, @{text rec_t}, and @{text size_t}.}
54537
f37354a894a3 fixed LaTeX missing }
blanchet
parents: 54494
diff changeset
  1003
f37354a894a3 fixed LaTeX missing }
blanchet
parents: 54494
diff changeset
  1004
\item \emph{The recursor @{text rec_t} has a different signature for nested
54185
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
  1005
recursive datatypes.} In the old package, nested recursion through non-functions
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
  1006
was internally reduced to mutual recursion. This reduction was visible in the
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
  1007
type of the recursor, used by \keyw{primrec}. Recursion through functions was
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
  1008
handled specially. In the new package, nested recursion (for functions and
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
  1009
non-functions) is handled in a more modular fashion. The old-style recursor can
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1010
be generated on demand using @{command primrec} if the recursion is via
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1011
new-style datatypes, as explained in
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1012
Section~\ref{sssec:primrec-nested-as-mutual-recursion}.
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1013
54287
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  1014
\item \emph{Accordingly, the induction rule is different for nested recursive
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  1015
datatypes.} Again, the old-style induction rule can be generated on demand using
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1016
@{command primrec} if the recursion is via new-style datatypes, as explained in
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1017
Section~\ref{sssec:primrec-nested-as-mutual-recursion}.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1018
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
  1019
\item \emph{The internal constructions are completely different.} Proof texts
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1020
that unfold the definition of constants introduced by \keyw{datatype} will be
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1021
difficult to port.
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1022
56644
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1023
\item \emph{Some theorems have different names.}
55641
5b466efedd2c renamed 'recs' and 'cases' theorems 'rec' and 'case' in old datatype package, for consistency with new package
blanchet
parents: 55531
diff changeset
  1024
For non-mutually recursive datatypes,
5b466efedd2c renamed 'recs' and 'cases' theorems 'rec' and 'case' in old datatype package, for consistency with new package
blanchet
parents: 55531
diff changeset
  1025
the alias @{text t.inducts} for @{text t.induct} is no longer generated.
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1026
For $m > 1$ mutually recursive datatypes,
53997
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1027
@{text "t\<^sub>1_\<dots>_t\<^sub>m.inducts(i)"} has been renamed
56644
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1028
@{text "t\<^sub>i.induct"} for each @{text "i \<in> {1, \<dots>, t}"}, and similarly the
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1029
collection @{text "t\<^sub>1_\<dots>_t\<^sub>m.size"} has been divided into @{text "t\<^sub>1.size"},
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1030
\ldots, @{text "t\<^sub>m.size"}.
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1031
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1032
\item \emph{The @{text t.simps} collection has been extended.}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1033
Previously available theorems are available at the same index.
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1034
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1035
\item \emph{Variables in generated properties have different names.} This is
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1036
rarely an issue, except in proof texts that refer to variable names in the
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1037
@{text "[where \<dots>]"} attribute. The solution is to use the more robust
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1038
@{text "[of \<dots>]"} syntax.
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1039
\end{itemize}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1040
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1041
In the other direction, there is currently no way to register old-style
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1042
datatypes as new-style datatypes. If the goal is to define new-style datatypes
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1043
with nested recursion through old-style datatypes, the old-style
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1044
datatypes can be registered as a BNF
55351
blanchet
parents: 55350
diff changeset
  1045
(Section~\ref{sec:introducing-bounded-natural-functors}). If the goal is
55468
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
  1046
to derive discriminators and selectors, this can be achieved using
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
  1047
@{command free_constructors}
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1048
(Section~\ref{sec:deriving-destructors-and-theorems-for-free-constructors}).
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1049
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1050
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1051
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1052
section {* Defining Recursive Functions
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1053
  \label{sec:defining-recursive-functions} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1054
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1055
text {*
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1056
Recursive functions over datatypes can be specified using the @{command primrec}
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1057
command, which supports primitive recursion, or using the more general
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1058
\keyw{fun} and \keyw{function} commands. Here, the focus is on
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1059
@{command primrec}; the other two commands are described in a separate
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1060
tutorial \cite{isabelle-function}.
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1061
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1062
%%% TODO: partial_function
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1063
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1064
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1065
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1066
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1067
  \label{ssec:primrec-introductory-examples} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1068
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1069
text {*
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1070
Primitive recursion is illustrated through concrete examples based on the
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1071
datatypes defined in Section~\ref{ssec:datatype-introductory-examples}. More
55075
b3d0a02a756d dissolved BNF session
blanchet
parents: 55073
diff changeset
  1072
examples can be found in the directory \verb|~~/src/HOL/BNF_Examples|.
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1073
*}
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1074
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1075
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1076
subsubsection {* Nonrecursive Types
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1077
  \label{sssec:primrec-nonrecursive-types} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1078
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1079
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1080
Primitive recursion removes one layer of constructors on the left-hand side in
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1081
each equation. For example:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1082
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1083
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1084
    primrec bool_of_trool :: "trool \<Rightarrow> bool" where
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1085
      "bool_of_trool Faalse \<longleftrightarrow> False" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1086
      "bool_of_trool Truue \<longleftrightarrow> True"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1087
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1088
text {* \blankline *}
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1089
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1090
    primrec the_list :: "'a option \<Rightarrow> 'a list" where
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1091
      "the_list None = []" |
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1092
      "the_list (Some a) = [a]"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1093
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1094
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1095
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1096
    primrec the_default :: "'a \<Rightarrow> 'a option \<Rightarrow> 'a" where
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1097
      "the_default d None = d" |
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1098
      "the_default _ (Some a) = a"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1099
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1100
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1101
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1102
    primrec mirrror :: "('a, 'b, 'c) triple \<Rightarrow> ('c, 'b, 'a) triple" where
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1103
      "mirrror (Triple a b c) = Triple c b a"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1104
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1105
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1106
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1107
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
  1108
some cases, which are then unspecified. Pattern matching on the left-hand side
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1109
is restricted to a single datatype, which must correspond to the same argument
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1110
in all equations.
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1111
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1112
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1113
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1114
subsubsection {* Simple Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1115
  \label{sssec:primrec-simple-recursion} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1116
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1117
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1118
For simple recursive types, recursive calls on a constructor argument are
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1119
allowed on the right-hand side:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1120
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1121
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1122
    primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1123
      "replicate Zero _ = []" |
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1124
      "replicate (Suc n) x = x # replicate n x"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1125
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1126
text {* \blankline *}
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1127
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1128
    primrec at :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" where
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1129
      "at (x # xs) j =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1130
         (case j of
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1131
            Zero \<Rightarrow> x
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1132
          | Suc j' \<Rightarrow> at xs j')"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1133
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1134
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1135
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1136
    primrec (*<*)(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
  1137
      "tfold _ (TNil y) = y" |
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1138
      "tfold f (TCons x xs) = f x (tfold f xs)"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1139
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1140
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1141
\noindent
54402
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1142
Pattern matching is only available for the argument on which the recursion takes
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1143
place. Fortunately, it is easy to generate pattern-maching equations using the
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1144
\keyw{simps\_of\_case} command provided by the theory
55290
3951ced4156c searchable underscores
blanchet
parents: 55254
diff changeset
  1145
\verb|~~/src/HOL/|\allowbreak\verb|Library/|\allowbreak\verb|Simps_Case_Conv|.
54402
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1146
*}
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1147
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1148
    simps_of_case at_simps: at.simps
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1149
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1150
text {*
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1151
This generates the lemma collection @{thm [source] at_simps}:
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1152
%
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1153
\[@{thm at_simps(1)[no_vars]}
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1154
  \qquad @{thm at_simps(2)[no_vars]}\]
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1155
%
54184
3fe73f3c84a2 more docs
blanchet
parents: 54183
diff changeset
  1156
The next example is defined using \keyw{fun} to escape the syntactic
55254
2bc951e6761a 'primitive' is not an adverb
blanchet
parents: 55129
diff changeset
  1157
restrictions imposed on primitively recursive functions. The
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
  1158
@{command datatype_compat} command is needed to register new-style datatypes
54184
3fe73f3c84a2 more docs
blanchet
parents: 54183
diff changeset
  1159
for use with \keyw{fun} and \keyw{function}
56644
efb39e0a89b0 updated docs
blanchet
parents: 56124
diff changeset
  1160
(Section~\ref{sssec:datatype-compat}):
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1161
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1162
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
  1163
    datatype_compat nat
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1164
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1165
text {* \blankline *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1166
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1167
    fun at_least_two :: "nat \<Rightarrow> bool" where
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1168
      "at_least_two (Suc (Suc _)) \<longleftrightarrow> True" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1169
      "at_least_two _ \<longleftrightarrow> False"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1170
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1171
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1172
subsubsection {* Mutual Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1173
  \label{sssec:primrec-mutual-recursion} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1174
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1175
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1176
The syntax for mutually recursive functions over mutually recursive datatypes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1177
is straightforward:
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1178
*}
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1179
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1180
    primrec
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1181
      nat_of_even_nat :: "even_nat \<Rightarrow> nat" and
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1182
      nat_of_odd_nat :: "odd_nat \<Rightarrow> nat"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1183
    where
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1184
      "nat_of_even_nat Even_Zero = Zero" |
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1185
      "nat_of_even_nat (Even_Suc n) = Suc (nat_of_odd_nat n)" |
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1186
      "nat_of_odd_nat (Odd_Suc n) = Suc (nat_of_even_nat n)"
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1187
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1188
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1189
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1190
    primrec
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1191
      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
  1192
      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
  1193
      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
  1194
    where
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1195
      "eval\<^sub>e \<gamma> \<xi> (Term t) = eval\<^sub>t \<gamma> \<xi> t" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1196
      "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
  1197
      "eval\<^sub>t \<gamma> \<xi> (Factor f) = eval\<^sub>f \<gamma> \<xi> f" |
52841
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1198
      "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
  1199
      "eval\<^sub>f \<gamma> _ (Const a) = \<gamma> a" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1200
      "eval\<^sub>f _ \<xi> (Var b) = \<xi> b" |
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1201
      "eval\<^sub>f \<gamma> \<xi> (Expr e) = eval\<^sub>e \<gamma> \<xi> e"
87a66bad0796 more (co)datatype documentation
blanchet
parents: 52840
diff changeset
  1202
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1203
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1204
\noindent
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1205
Mutual recursion is possible within a single type, using \keyw{fun}:
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1206
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1207
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1208
    fun
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1209
      even :: "nat \<Rightarrow> bool" and
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1210
      odd :: "nat \<Rightarrow> bool"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1211
    where
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1212
      "even Zero = True" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1213
      "even (Suc n) = odd n" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1214
      "odd Zero = False" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1215
      "odd (Suc n) = even n"
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1216
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1217
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1218
subsubsection {* Nested Recursion
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1219
  \label{sssec:primrec-nested-recursion} *}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1220
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1221
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1222
In a departure from the old datatype package, nested recursion is normally
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1223
handled via the map functions of the nesting type constructors. For example,
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1224
recursive calls are lifted to lists using @{const map}:
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1225
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1226
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1227
(*<*)
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1228
    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
  1229
(*>*)
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1230
    primrec at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \<Rightarrow> nat list \<Rightarrow> 'a" where
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1231
      "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1232
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1233
            [] \<Rightarrow> a
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  1234
          | 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
  1235
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1236
text {*
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1237
\noindent
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1238
The next example features recursion through the @{text option} type. Although
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1239
@{text option} is not a new-style datatype, it is registered as a BNF with the
54491
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
  1240
map function @{const map_option}:
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1241
*}
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1242
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1243
    primrec (*<*)(in early) (*>*)sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" where
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1244
      "sum_btree (BNode a lt rt) =
54491
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
  1245
         a + the_default 0 (map_option sum_btree lt) +
27966e17d075 case_if -> case_eq_if + docs
blanchet
parents: 54422
diff changeset
  1246
           the_default 0 (map_option sum_btree rt)"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1247
53136
98a2c33d5d1b ideas for (co)datatype docs
blanchet
parents: 53028
diff changeset
  1248
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1249
\noindent
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1250
The same principle applies for arbitrary type constructors through which
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1251
recursion is possible. Notably, the map function for the function type
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1252
(@{text \<Rightarrow>}) is simply composition (@{text "op \<circ>"}):
53136
98a2c33d5d1b ideas for (co)datatype docs
blanchet
parents: 53028
diff changeset
  1253
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1254
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1255
    primrec (*<*)(in early) (*>*)relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1256
      "relabel_ft f (FTLeaf x) = FTLeaf (f x)" |
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1257
      "relabel_ft f (FTNode g) = FTNode (relabel_ft f \<circ> g)"
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1258
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1259
text {*
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1260
\noindent
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1261
For convenience, recursion through functions can also be expressed using
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1262
$\lambda$-abstractions and function application rather than through composition.
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1263
For example:
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1264
*}
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1265
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1266
    primrec relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1267
      "relabel_ft f (FTLeaf x) = FTLeaf (f x)" |
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1268
      "relabel_ft f (FTNode g) = FTNode (\<lambda>x. relabel_ft f (g x))"
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1269
54183
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1270
text {* \blankline *}
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1271
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1272
    primrec subtree_ft :: "'a \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
54183
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1273
      "subtree_ft x (FTNode g) = g x"
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1274
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1275
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1276
\noindent
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1277
For recursion through curried $n$-ary functions, $n$ applications of
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1278
@{term "op \<circ>"} are necessary. The examples below illustrate the case where
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1279
$n = 2$:
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1280
*}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1281
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1282
    datatype_new 'a ftree2 = FTLeaf2 'a | FTNode2 "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2"
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1283
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1284
text {* \blankline *}
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1285
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1286
    primrec (*<*)(in early) (*>*)relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1287
      "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" |
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1288
      "relabel_ft2 f (FTNode2 g) = FTNode2 (op \<circ> (op \<circ> (relabel_ft2 f)) g)"
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1289
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1290
text {* \blankline *}
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1291
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1292
    primrec relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1293
      "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" |
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  1294
      "relabel_ft2 f (FTNode2 g) = FTNode2 (\<lambda>x y. relabel_ft2 f (g x y))"
54031
a3281fbe6856 more (co)data docs
blanchet
parents: 54023
diff changeset
  1295
54183
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1296
text {* \blankline *}
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1297
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1298
    primrec subtree_ft2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
54183
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1299
      "subtree_ft2 x y (FTNode2 g) = g x y"
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1300
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1301
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1302
subsubsection {* Nested-as-Mutual Recursion
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1303
  \label{sssec:primrec-nested-as-mutual-recursion} *}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1304
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1305
(*<*)
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1306
    locale n2m begin
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1307
(*>*)
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1308
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1309
text {*
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1310
For compatibility with the old package, but also because it is sometimes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1311
convenient in its own right, it is possible to treat nested recursive datatypes
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1312
as mutually recursive ones if the recursion takes place though new-style
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1313
datatypes. For example:
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1314
*}
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1315
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1316
    primrec
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1317
      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
  1318
      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
  1319
    where
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1320
      "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1321
         (case js of
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1322
            [] \<Rightarrow> a
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1323
          | j # js' \<Rightarrow> ats\<^sub>f\<^sub>f ts j js')" |
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1324
      "ats\<^sub>f\<^sub>f (t # ts) j =
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1325
         (case j of
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1326
            Zero \<Rightarrow> at\<^sub>f\<^sub>f t
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1327
          | Suc j' \<Rightarrow> ats\<^sub>f\<^sub>f ts j')"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1328
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1329
text {*
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1330
\noindent
54287
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  1331
Appropriate induction rules are generated as
54031
a3281fbe6856 more (co)data docs
blanchet
parents: 54023
diff changeset
  1332
@{thm [source] at\<^sub>f\<^sub>f.induct},
a3281fbe6856 more (co)data docs
blanchet
parents: 54023
diff changeset
  1333
@{thm [source] ats\<^sub>f\<^sub>f.induct}, and
54287
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  1334
@{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  1335
induction rules and the underlying recursors are generated on a per-need basis
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  1336
and are kept in a cache to speed up subsequent definitions.
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1337
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1338
Here is a second example:
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1339
*}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1340
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1341
    primrec
53330
77da8d3c46e0 fixed docs w.r.t. availability of "primrec_new" and friends
blanchet
parents: 53262
diff changeset
  1342
      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
  1343
      sum_btree_option :: "'a btree option \<Rightarrow> 'a"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1344
    where
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1345
      "sum_btree (BNode a lt rt) =
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1346
         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
  1347
      "sum_btree_option None = 0" |
53025
c820c9e9e8f4 more work on (co)datatype docs
blanchet
parents: 53018
diff changeset
  1348
      "sum_btree_option (Some t) = sum_btree t"
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1349
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1350
text {*
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1351
%  * can pretend a nested type is mutually recursive (if purely inductive)
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1352
%  * avoids the higher-order map
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1353
%  * e.g.
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1354
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1355
%  * this can always be avoided;
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1356
%     * e.g. in our previous example, we first mapped the recursive
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1357
%       calls, then we used a generic at function to retrieve the result
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1358
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1359
%  * 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
  1360
%    just like there's no rule when to use fold and when to use
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1361
%    primrec
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1362
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1363
%  * higher-order approach, considering nesting as nesting, is more
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1364
%    compositional -- e.g. we saw how we could reuse an existing polymorphic
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1365
%    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
  1366
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1367
%  * but:
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1368
%     * is perhaps less intuitive, because it requires higher-order thinking
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1369
%     * may seem inefficient, and indeed with the code generator the
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1370
%       mutually recursive version might be nicer
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1371
%     * is somewhat indirect -- must apply a map first, then compute a result
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1372
%       (cannot mix)
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1373
%     * 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
  1374
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1375
%  * impact on automation unclear
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1376
%
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1377
*}
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1378
(*<*)
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1379
    end
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1380
(*>*)
52843
ea95702328cf more (co)datatype docs
blanchet
parents: 52841
diff changeset
  1381
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1382
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1383
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1384
  \label{ssec:primrec-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1385
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1386
56123
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1387
subsubsection {* \keyw{primrec}
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1388
  \label{sssec:primrec-new} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1389
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1390
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1391
\begin{matharray}{rcl}
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1392
  @{command_def "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"}
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1393
\end{matharray}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1394
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1395
@{rail \<open>
56123
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1396
  @@{command primrec} target? @{syntax pr_option}? fixes \<newline>
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1397
  @'where' (@{syntax pr_equation} + '|')
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1398
  ;
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1399
  @{syntax_def pr_option}: '(' 'nonexhaustive' ')'
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1400
  ;
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1401
  @{syntax_def pr_equation}: thmdecl? prop
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1402
\<close>}
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1403
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1404
\medskip
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1405
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1406
\noindent
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1407
The @{command primrec} command introduces a set of mutually recursive functions
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1408
over datatypes.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1409
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1410
The syntactic entity \synt{target} can be used to specify a local context,
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  1411
\synt{fixes} denotes a list of names with optional type signatures,
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  1412
\synt{thmdecl} denotes an optional name for the formula that follows, and
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  1413
\synt{prop} denotes a HOL proposition \cite{isabelle-isar-ref}.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1414
56124
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
  1415
The optional target is optionally followed by the following option:
56123
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1416
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1417
\begin{itemize}
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1418
\setlength{\itemsep}{0pt}
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1419
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1420
\item
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1421
The @{text "nonexhaustive"} option indicates that the functions are not
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1422
necessarily specified for all constructors. It can be used to suppress the
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1423
warning that is normally emitted when some constructors are missing.
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1424
\end{itemize}
a27859b0ef7d document the new 'nonexhaustive' option (cf. 52e8f110fec3)
blanchet
parents: 55945
diff changeset
  1425
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1426
%%% TODO: elaborate on format of the equations
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  1427
%%% TODO: elaborate on mutual and nested-to-mutual
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1428
*}
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1429
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1430
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1431
(*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1432
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1433
  \label{ssec:primrec-generated-theorems} *}
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1434
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1435
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1436
%  * synthesized nonrecursive definition
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1437
%  * user specification is rederived from it, exactly as entered
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1438
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1439
%  * induct
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1440
%    * mutualized
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1441
%    * without some needless induction hypotheses if not used
55871
a28817253b31 removed (co)iterators from documentation
blanchet
parents: 55705
diff changeset
  1442
%  * rec
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1443
%    * mutualized
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1444
*}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1445
*)
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1446
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1447
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  1448
subsection {* Recursive Default Values for Selectors
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1449
  \label{ssec:primrec-recursive-default-values-for-selectors} *}
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1450
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1451
text {*
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1452
A datatype selector @{text un_D} can have a default value for each constructor
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1453
on which it is not otherwise specified. Occasionally, it is useful to have the
55354
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1454
default value be defined recursively. This leads to a chicken-and-egg
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1455
situation, because the datatype is not introduced yet at the moment when the
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1456
selectors are introduced. Of course, we can always define the selectors
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1457
manually afterward, but we then have to state and prove all the characteristic
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1458
theorems ourselves instead of letting the package do it.
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1459
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1460
Fortunately, there is a workaround that relies on overloading to relieve us
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1461
from the tedium of manual derivations:
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1462
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1463
\begin{enumerate}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1464
\setlength{\itemsep}{0pt}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1465
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1466
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1467
Introduce a fully unspecified constant @{text "un_D\<^sub>0 \<Colon> 'a"} using
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1468
@{keyword consts}.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1469
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1470
\item
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1471
Define the datatype, specifying @{text "un_D\<^sub>0"} as the selector's default
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1472
value.
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1473
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1474
\item
53535
d0c163c6c725 more (co)data docs
blanchet
parents: 53534
diff changeset
  1475
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
  1476
datatype using the \keyw{overloading} command.
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1477
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1478
\item
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1479
Derive the desired equation on @{text un_D} from the characteristic equations
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1480
for @{text "un_D\<^sub>0"}.
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1481
\end{enumerate}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1482
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1483
\noindent
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1484
The following example illustrates this procedure:
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1485
*}
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1486
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1487
(*<*)
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1488
    hide_const termi
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1489
(*>*)
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1490
    consts termi\<^sub>0 :: 'a
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1491
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1492
text {* \blankline *}
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1493
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
  1494
    datatype_new ('a, 'b) tlist =
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1495
      TNil (termi: 'b)
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1496
    | TCons (thd: 'a) (ttl: "('a, 'b) tlist")
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1497
    where
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1498
      "ttl (TNil y) = TNil y"
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1499
    | "termi (TCons _ xs) = termi\<^sub>0 xs"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1500
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1501
text {* \blankline *}
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1502
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1503
    overloading
53491
2479b39d9d09 more docs
blanchet
parents: 53335
diff changeset
  1504
      termi\<^sub>0 \<equiv> "termi\<^sub>0 \<Colon> ('a, 'b) tlist \<Rightarrow> 'b"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1505
    begin
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1506
    primrec termi\<^sub>0 :: "('a, 'b) tlist \<Rightarrow> 'b" where
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1507
      "termi\<^sub>0 (TNil y) = y" |
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1508
      "termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1509
    end
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1510
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1511
text {* \blankline *}
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1512
55354
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1513
    lemma termi_TCons[simp]: "termi (TCons x xs) = termi xs"
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1514
    by (cases xs) auto
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1515
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1516
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1517
subsection {* Compatibility Issues
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1518
  \label{ssec:primrec-compatibility-issues} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1519
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1520
text {*
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1521
The command @{command primrec}'s behavior on new-style datatypes has been
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1522
designed to be highly compatible with that for old-style datatypes, to ease
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1523
migration. There is nonetheless at least one incompatibility that may arise when
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55474
diff changeset
  1524
porting to the new package:
53997
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1525
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1526
\begin{itemize}
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1527
\setlength{\itemsep}{0pt}
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1528
54185
3fe3b5d33e41 more docs
blanchet
parents: 54184
diff changeset
  1529
\item \emph{Some theorems have different names.}
53997
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1530
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
  1531
@{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
  1532
subcollections @{text "f\<^sub>i.simps"}.
53997
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1533
\end{itemize}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1534
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1535
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1536
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1537
section {* Defining Codatatypes
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1538
  \label{sec:defining-codatatypes} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1539
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1540
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1541
Codatatypes can be specified using the @{command codatatype} command. The
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1542
command is first illustrated through concrete examples featuring different
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1543
flavors of corecursion. More examples can be found in the directory
53997
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1544
\verb|~~/src/HOL/|\allowbreak\verb|BNF/Examples|. The
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1545
\emph{Archive of Formal Proofs} also includes some useful codatatypes, notably
8ff43f638da2 more "primrec_new" documentation
blanchet
parents: 53917
diff changeset
  1546
for lazy lists \cite{lochbihler-2010}.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1547
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  1548
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1549
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1550
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1551
  \label{ssec:codatatype-introductory-examples} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  1552
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1553
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1554
subsubsection {* Simple Corecursion
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1555
  \label{sssec:codatatype-simple-corecursion} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1556
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1557
text {*
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
  1558
Noncorecursive codatatypes coincide with the corresponding datatypes, so they
56947
01ab2e94a713 tuned docs
blanchet
parents: 56942
diff changeset
  1559
are rarely used in practice. \emph{Corecursive codatatypes} have the same syntax
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1560
as recursive datatypes, except for the command name. For example, here is the
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1561
definition of lazy lists:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1562
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1563
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1564
    codatatype (lset: 'a) llist =
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1565
      lnull: LNil
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1566
    | LCons (lhd: 'a) (ltl: "'a llist")
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1567
    for
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1568
      map: lmap
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1569
      rel: llist_all2
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1570
    where
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  1571
      "ltl LNil = LNil"
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1572
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1573
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1574
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1575
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
  1576
@{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
  1577
infinite streams:
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1578
*}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1579
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1580
    codatatype (sset: 'a) stream =
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1581
      SCons (shd: 'a) (stl: "'a stream")
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1582
    for
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1583
      map: smap
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  1584
      rel: stream_all2
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1585
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1586
text {*
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1587
\noindent
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1588
Another interesting type that can
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1589
be defined as a codatatype is that of the extended natural numbers:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1590
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1591
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1592
    codatatype enat = EZero | ESuc enat
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1593
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1594
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1595
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1596
This type has exactly one infinite element, @{text "ESuc (ESuc (ESuc (\<dots>)))"},
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1597
that represents $\infty$. In addition, it has finite values of the form
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1598
@{text "ESuc (\<dots> (ESuc EZero)\<dots>)"}.
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1599
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1600
Here is an example with many constructors:
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1601
*}
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1602
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1603
    codatatype 'a process =
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1604
      Fail
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1605
    | Skip (cont: "'a process")
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1606
    | Action (prefix: 'a) (cont: "'a process")
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1607
    | Choice (left: "'a process") (right: "'a process")
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1608
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  1609
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1610
\noindent
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  1611
Notice that the @{const cont} selector is associated with both @{const Skip}
54146
97f69d44f732 doc fixes suggested by Andreas L.
blanchet
parents: 54072
diff changeset
  1612
and @{const Action}.
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  1613
*}
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  1614
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1615
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1616
subsubsection {* Mutual Corecursion
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1617
  \label{sssec:codatatype-mutual-corecursion} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1618
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1619
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1620
\noindent
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1621
The example below introduces a pair of \emph{mutually corecursive} types:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1622
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1624
    codatatype even_enat = Even_EZero | Even_ESuc odd_enat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1625
    and odd_enat = Odd_ESuc even_enat
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1626
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1627
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1628
subsubsection {* Nested Corecursion
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1629
  \label{sssec:codatatype-nested-corecursion} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1630
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1631
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1632
\noindent
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1633
The next examples feature \emph{nested corecursion}:
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1634
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1635
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1636
    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
  1637
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1638
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1639
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1640
    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
  1641
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1642
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1643
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  1644
    codatatype 'a sm = SM (accept: bool) (trans: "'a \<Rightarrow> 'a sm")
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1645
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1646
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1647
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1648
  \label{ssec:codatatype-command-syntax} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1649
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1650
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1651
subsubsection {* \keyw{codatatype}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  1652
  \label{sssec:codatatype} *}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  1653
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1654
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1655
\begin{matharray}{rcl}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1656
  @{command_def "codatatype"} & : & @{text "local_theory \<rightarrow> local_theory"}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1657
\end{matharray}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1658
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1659
@{rail \<open>
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 54958
diff changeset
  1660
  @@{command codatatype} target? \<newline>
55472
990651bfc65b updated docs to reflect the new 'free_constructors' syntax
blanchet
parents: 55468
diff changeset
  1661
    (@{syntax dt_name} '=' (@{syntax dt_ctor} + '|') + @'and')
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  1662
\<close>}
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1663
55351
blanchet
parents: 55350
diff changeset
  1664
\medskip
blanchet
parents: 55350
diff changeset
  1665
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1666
\noindent
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1667
Definitions of codatatypes have almost exactly the same syntax as for datatypes
57094
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
  1668
(Section~\ref{ssec:datatype-command-syntax}). The @{text "discs_sels"} option
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
  1669
is superfluous because discriminators and selectors are always generated for
589ec121ce1a don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
blanchet
parents: 57092
diff changeset
  1670
codatatypes.
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1671
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1672
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1673
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1674
subsection {* Generated Constants
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1675
  \label{ssec:codatatype-generated-constants} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1676
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1677
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1678
Given a codatatype @{text "('a\<^sub>1, \<dots>, 'a\<^sub>m) t"}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1679
with $m > 0$ live type variables and $n$ constructors @{text "t.C\<^sub>1"},
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1680
\ldots, @{text "t.C\<^sub>n"}, the same auxiliary constants are generated as for
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1681
datatypes (Section~\ref{ssec:datatype-generated-constants}), except that the
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1682
recursor is replaced by a dual concept and no size function is produced:
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1683
55354
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1684
\medskip
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1685
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1686
\begin{tabular}{@ {}ll@ {}}
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1687
Corecursor: &
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1688
  @{text t.corec_t}
55354
6ca9df01ac8c more docs
blanchet
parents: 55353
diff changeset
  1689
\end{tabular}
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1690
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1691
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1692
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1693
subsection {* Generated Theorems
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1694
  \label{ssec:codatatype-generated-theorems} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1695
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1696
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1697
The characteristic theorems generated by @{command codatatype} are grouped in
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1698
three broad categories:
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1699
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1700
\begin{itemize}
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1701
\setlength{\itemsep}{0pt}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1702
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1703
\item The \emph{free constructor theorems}
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1704
(Section~\ref{sssec:free-constructor-theorems}) are properties of the
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1705
constructors and destructors that can be derived for any freely generated type.
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1706
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1707
\item The \emph{functorial theorems}
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1708
(Section~\ref{sssec:functorial-theorems}) are properties of datatypes related to
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1709
their BNF nature.
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1710
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1711
\item The \emph{coinductive theorems} (Section~\ref{sssec:coinductive-theorems})
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1712
are properties of datatypes related to their coinductive nature.
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1713
\end{itemize}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1714
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1715
\noindent
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  1716
The first two categories are exactly as for datatypes.
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1717
*}
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1718
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1719
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1720
subsubsection {* Coinductive Theorems
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1721
  \label{sssec:coinductive-theorems} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1722
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1723
text {*
54031
a3281fbe6856 more (co)data docs
blanchet
parents: 54023
diff changeset
  1724
The coinductive theorems are listed below for @{typ "'a llist"}:
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1725
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1726
\begin{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1727
\begin{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1728
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1729
\item[\begin{tabular}{@ {}l@ {}}
57304
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1730
  @{text "t."}\hthm{coinduct} @{text "[consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1731
  \phantom{@{text "t."}\hthm{coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots>
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1732
  D\<^sub>n, coinduct t]"}\rm:
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1733
\end{tabular}] ~ \\
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1734
@{thm llist.coinduct[no_vars]}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1735
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1736
\item[\begin{tabular}{@ {}l@ {}}
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1737
  @{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
  1738
  \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
  1739
\end{tabular}] ~ \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1740
@{thm llist.strong_coinduct[no_vars]}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1741
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1742
\item[\begin{tabular}{@ {}l@ {}}
57304
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1743
  @{text "t."}\hthm{rel_coinduct} @{text "[consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1744
  \phantom{@{text "t."}\hthm{rel_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots>
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1745
  D\<^sub>n, coinduct pred]"}\rm:
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1746
\end{tabular}] ~ \\
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1747
@{thm llist.rel_coinduct[no_vars]}
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1748
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1749
\item[\begin{tabular}{@ {}l@ {}}
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1750
  @{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
  1751
  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{strong\_coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
57304
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1752
  \phantom{@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{strong\_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm: \\
d2061dc953a4 document property 'rel_coinduct'
desharna
parents: 57206
diff changeset
  1753
  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{rel_coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m, case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"} \\
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1754
\end{tabular}] ~ \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1755
Given $m > 1$ mutually corecursive codatatypes, these coinduction rules can be
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1756
used to prove $m$ properties simultaneously.
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1757
54031
a3281fbe6856 more (co)data docs
blanchet
parents: 54023
diff changeset
  1758
\item[@{text "t."}\hthm{corec}\rm:] ~ \\
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1759
@{thm llist.corec(1)[no_vars]} \\
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1760
@{thm llist.corec(2)[no_vars]}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1761
53703
0c565fec9c78 updated docs
blanchet
parents: 53694
diff changeset
  1762
\item[@{text "t."}\hthm{disc\_corec}\rm:] ~ \\
53643
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1763
@{thm llist.disc_corec(1)[no_vars]} \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1764
@{thm llist.disc_corec(2)[no_vars]}
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1765
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1766
\item[@{text "t."}\hthm{disc\_corec\_iff} @{text "[simp]"}\rm:] ~ \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1767
@{thm llist.disc_corec_iff(1)[no_vars]} \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1768
@{thm llist.disc_corec_iff(2)[no_vars]}
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1769
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1770
\item[@{text "t."}\hthm{sel\_corec} @{text "[simp]"}\rm:] ~ \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1771
@{thm llist.sel_corec(1)[no_vars]} \\
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1772
@{thm llist.sel_corec(2)[no_vars]}
673cb2c9d695 more (co)data docs
blanchet
parents: 53642
diff changeset
  1773
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1774
\end{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1775
\end{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1776
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1777
\noindent
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  1778
For convenience, @{command codatatype} also provides the following collection:
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1779
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1780
\begin{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1781
\begin{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1782
55896
c78575827f38 minor doc fix
blanchet
parents: 55871
diff changeset
  1783
\item[@{text "t."}\hthm{simps} = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.disc_corec_iff}] @{text t.sel_corec} ~ \\
c78575827f38 minor doc fix
blanchet
parents: 55871
diff changeset
  1784
@{text t.map} @{text t.rel_inject} @{text t.rel_distinct} @{text t.set}
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1785
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1786
\end{description}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1787
\end{indentblock}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  1788
*}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1789
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1790
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  1791
section {* Defining Corecursive Functions
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1792
  \label{sec:defining-corecursive-functions} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1793
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1794
text {*
54183
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1795
Corecursive functions can be specified using the @{command primcorec} and
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1796
\keyw{prim\-corec\-ursive} commands, which support primitive corecursion, or
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1797
using the more general \keyw{partial\_function} command. Here, the focus is on
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1798
the first two. More examples can be found in the directory
55075
b3d0a02a756d dissolved BNF session
blanchet
parents: 55073
diff changeset
  1799
\verb|~~/src/HOL/BNF_Examples|.
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1800
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1801
Whereas recursive functions consume datatypes one constructor at a time,
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1802
corecursive functions construct codatatypes one constructor at a time.
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1803
Partly reflecting a lack of agreement among proponents of coalgebraic methods,
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1804
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
  1805
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1806
\begin{itemize}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1807
\setlength{\itemsep}{0pt}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1808
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1809
\abovedisplayskip=.5\abovedisplayskip
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1810
\belowdisplayskip=.5\belowdisplayskip
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1811
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1812
\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
  1813
\[@{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
  1814
equations of the form
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1815
\[@{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
  1816
This style is popular in the coalgebraic literature.
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1817
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1818
\item The \emph{constructor view} specifies $f$ by equations of the form
54183
8a6a49385122 more docs
blanchet
parents: 54182
diff changeset
  1819
\[@{text "\<dots> \<Longrightarrow> f x\<^sub>1 \<dots> x\<^sub>n = C\<^sub>j \<dots>"}\]
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1820
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
  1821
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1822
\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
  1823
\[@{text "f x\<^sub>1 \<dots> x\<^sub>n = \<dots>"}\]
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1824
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
  1825
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
  1826
style.
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1827
\end{itemize}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1828
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
  1829
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
  1830
characteristic theorems for all three styles are generated.
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1831
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1832
%%% TODO: partial_function? E.g. for defining tail recursive function on lazy
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  1833
%%% lists (cf. terminal0 in TLList.thy)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1834
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1835
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  1836
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1837
subsection {* Introductory Examples
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  1838
  \label{ssec:primcorec-introductory-examples} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1839
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1840
text {*
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1841
Primitive corecursion is illustrated through concrete examples based on the
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1842
codatatypes defined in Section~\ref{ssec:codatatype-introductory-examples}. More
55075
b3d0a02a756d dissolved BNF session
blanchet
parents: 55073
diff changeset
  1843
examples can be found in the directory \verb|~~/src/HOL/BNF_Examples|. The code
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1844
view is favored in the examples below. Sections
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1845
\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
  1846
present the same examples expressed using the constructor and destructor views.
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1847
*}
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1848
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1849
subsubsection {* Simple Corecursion
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1850
  \label{sssec:primcorec-simple-corecursion} *}
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1851
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1852
text {*
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1853
Following the code view, corecursive calls are allowed on the right-hand side as
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1854
long as they occur under a constructor, which itself appears either directly to
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1855
the right of the equal sign or in a conditional expression:
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1856
*}
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1857
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1858
    primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1859
      "literate g x = LCons x (literate g (g x))"
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1860
53677
b51ebeda414d more (co)data docs
blanchet
parents: 53675
diff changeset
  1861
text {* \blankline *}
b51ebeda414d more (co)data docs
blanchet
parents: 53675
diff changeset
  1862
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1863
    primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1864
      "siterate g x = SCons x (siterate g (g x))"
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1865
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1866
text {*
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1867
\noindent
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1868
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
  1869
\emph{productive}. The above functions compute the infinite lazy list or stream
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1870
@{text "[x, g x, g (g x), \<dots>]"}. Productivity guarantees that prefixes
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1871
@{text "[x, g x, g (g x), \<dots>, (g ^^ k) x]"} of arbitrary finite length
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1872
@{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
  1873
times.
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1874
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1875
Corecursive functions construct codatatype values, but nothing prevents them
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
  1876
from also consuming such values. The following function drops every second
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1877
element in a stream:
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1878
*}
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1879
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1880
    primcorec every_snd :: "'a stream \<Rightarrow> 'a stream" where
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1881
      "every_snd s = SCons (shd s) (stl (stl s))"
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1882
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1883
text {*
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1884
\noindent
56124
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
  1885
Constructs such as @{text "let"}--@{text "in"}, @{text
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
  1886
"if"}--@{text "then"}--@{text "else"}, and @{text "case"}--@{text "of"} may
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1887
appear around constructors that guard corecursive calls:
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1888
*}
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1889
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1890
    primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1891
      "lappend xs ys =
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1892
         (case xs of
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1893
            LNil \<Rightarrow> ys
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1894
          | LCons x xs' \<Rightarrow> LCons x (lappend xs' ys))"
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1895
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1896
text {*
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1897
\noindent
54402
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1898
Pattern matching is not supported by @{command primcorec}. Fortunately, it is
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1899
easy to generate pattern-maching equations using the \keyw{simps\_of\_case}
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1900
command provided by the theory \verb|~~/src/HOL/Library/Simps_Case_Conv|.
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1901
*}
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1902
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1903
    simps_of_case lappend_simps: lappend.code
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1904
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1905
text {*
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1906
This generates the lemma collection @{thm [source] lappend_simps}:
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1907
%
55355
b5b64d9d1002 more docs
blanchet
parents: 55354
diff changeset
  1908
\begin{gather*%
b5b64d9d1002 more docs
blanchet
parents: 55354
diff changeset
  1909
}
b5b64d9d1002 more docs
blanchet
parents: 55354
diff changeset
  1910
  @{thm lappend_simps(1)[no_vars]} \\
b5b64d9d1002 more docs
blanchet
parents: 55354
diff changeset
  1911
  @{thm lappend_simps(2)[no_vars]}
b5b64d9d1002 more docs
blanchet
parents: 55354
diff changeset
  1912
\end{gather*%
b5b64d9d1002 more docs
blanchet
parents: 55354
diff changeset
  1913
}
54402
5d1b7ee6070e document idiomatic use of 'simps_of_case'
blanchet
parents: 54386
diff changeset
  1914
%
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1915
Corecursion is useful to specify not only functions but also infinite objects:
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1916
*}
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1917
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1918
    primcorec infty :: enat where
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1919
      "infty = ESuc infty"
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1920
53646
ac6e0a28489f more (co)data docs
blanchet
parents: 53644
diff changeset
  1921
text {*
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1922
\noindent
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1923
The example below constructs a pseudorandom process value. It takes a stream of
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1924
actions (@{text s}), a pseudorandom function generator (@{text f}), and a
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1925
pseudorandom seed (@{text n}):
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1926
*}
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1927
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1928
    primcorec
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1929
      random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1930
    where
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1931
      "random_process s f n =
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1932
         (if n mod 4 = 0 then
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1933
            Fail
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1934
          else if n mod 4 = 1 then
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1935
            Skip (random_process s f (f n))
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1936
          else if n mod 4 = 2 then
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1937
            Action (shd s) (random_process (stl s) f (f n))
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1938
          else
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1939
            Choice (random_process (every_snd s) (f \<circ> f) (f n))
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1940
              (random_process (every_snd (stl s)) (f \<circ> f) (f (f n))))"
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1941
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1942
text {*
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1943
\noindent
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1944
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
  1945
sequentially. This is visible in the generated theorems. The constructor and
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1946
destructor views offer nonsequential alternatives.
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1947
*}
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  1948
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1949
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1950
subsubsection {* Mutual Corecursion
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1951
  \label{sssec:primcorec-mutual-corecursion} *}
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1952
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1953
text {*
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1954
The syntax for mutually corecursive functions over mutually corecursive
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  1955
datatypes is unsurprising:
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1956
*}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1957
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1958
    primcorec
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1959
      even_infty :: even_enat and
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1960
      odd_infty :: odd_enat
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1961
    where
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1962
      "even_infty = Even_ESuc odd_infty" |
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1963
      "odd_infty = Odd_ESuc even_infty"
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1964
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1965
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1966
subsubsection {* Nested Corecursion
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1967
  \label{sssec:primcorec-nested-corecursion} *}
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1968
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1969
text {*
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1970
The next pair of examples generalize the @{const literate} and @{const siterate}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1971
functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1972
infinite trees in which subnodes are organized either as a lazy list (@{text
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1973
tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}). They rely on the map functions of
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1974
the nesting type constructors to lift the corecursive calls:
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1975
*}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  1976
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1977
    primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1978
      "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i g) (g x))"
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1979
53677
b51ebeda414d more (co)data docs
blanchet
parents: 53675
diff changeset
  1980
text {* \blankline *}
b51ebeda414d more (co)data docs
blanchet
parents: 53675
diff changeset
  1981
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  1982
    primcorec iterate\<^sub>i\<^sub>s :: "('a \<Rightarrow> 'a fset) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>s" where
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1983
      "iterate\<^sub>i\<^sub>s g x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s g) (g x))"
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  1984
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  1985
text {*
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  1986
\noindent
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1987
Both examples follow the usual format for constructor arguments associated
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1988
with nested recursive occurrences of the datatype. Consider
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1989
@{const iterate\<^sub>i\<^sub>i}. The term @{term "g x"} constructs an @{typ "'a llist"}
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1990
value, which is turned into an @{typ "'a tree\<^sub>i\<^sub>i llist"} value using
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1991
@{const lmap}.
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1992
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1993
This format may sometimes feel artificial. The following function constructs
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1994
a tree with a single, infinite branch from a stream:
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1995
*}
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1996
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1997
    primcorec tree\<^sub>i\<^sub>i_of_stream :: "'a stream \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1998
      "tree\<^sub>i\<^sub>i_of_stream s =
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  1999
         Node\<^sub>i\<^sub>i (shd s) (lmap tree\<^sub>i\<^sub>i_of_stream (LCons (stl s) LNil))"
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2000
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2001
text {*
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2002
\noindent
54955
cf8d429dc24e reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
blanchet
parents: 54832
diff changeset
  2003
A more natural syntax, also supported by Isabelle, is to move corecursive calls
cf8d429dc24e reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
blanchet
parents: 54832
diff changeset
  2004
under constructors:
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2005
*}
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2006
54955
cf8d429dc24e reintroduce recursive calls under constructors, taken out in 8dd0e0316881 mainly and in subsequent changes
blanchet
parents: 54832
diff changeset
  2007
    primcorec (*<*)(in late) (*>*)tree\<^sub>i\<^sub>i_of_stream :: "'a stream \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2008
      "tree\<^sub>i\<^sub>i_of_stream s =
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2009
         Node\<^sub>i\<^sub>i (shd s) (LCons (tree\<^sub>i\<^sub>i_of_stream (stl s)) LNil)"
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2010
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2011
text {*
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2012
The next example illustrates corecursion through functions, which is a bit
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2013
special. Deterministic finite automata (DFAs) are traditionally defined as
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2014
5-tuples @{text "(Q, \<Sigma>, \<delta>, q\<^sub>0, F)"}, where @{text Q} is a finite set of states,
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  2015
@{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
  2016
is an initial state, and @{text F} is a set of final states. The following
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2017
function translates a DFA into a state machine:
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  2018
*}
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  2019
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2020
    primcorec (*<*)(in early) (*>*)sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a sm" where
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2021
      "sm_of_dfa \<delta> F q = SM (q \<in> F) (sm_of_dfa \<delta> F \<circ> \<delta> q)"
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  2022
53751
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2023
text {*
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2024
\noindent
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2025
The map function for the function type (@{text \<Rightarrow>}) is composition
54181
66edcd52daeb updated doc
blanchet
parents: 54152
diff changeset
  2026
(@{text "op \<circ>"}). For convenience, corecursion through functions can
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2027
also be expressed using $\lambda$-abstractions and function application rather
54031
a3281fbe6856 more (co)data docs
blanchet
parents: 54023
diff changeset
  2028
than through composition. For example:
53751
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2029
*}
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2030
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2031
    primcorec sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a sm" where
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2032
      "sm_of_dfa \<delta> F q = SM (q \<in> F) (\<lambda>a. sm_of_dfa \<delta> F (\<delta> q a))"
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2033
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2034
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2035
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2036
    primcorec empty_sm :: "'a sm" where
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2037
      "empty_sm = SM False (\<lambda>_. empty_sm)"
53751
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2038
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2039
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2040
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2041
    primcorec not_sm :: "'a sm \<Rightarrow> 'a sm" where
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2042
      "not_sm M = SM (\<not> accept M) (\<lambda>a. not_sm (trans M a))"
53751
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
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
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2046
    primcorec or_sm :: "'a sm \<Rightarrow> 'a sm \<Rightarrow> 'a sm" where
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2047
      "or_sm M N =
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2048
         SM (accept M \<or> accept N) (\<lambda>a. or_sm (trans M a) (trans N a))"
53751
7a994dc08cea added primcorec examples with lambdas
blanchet
parents: 53750
diff changeset
  2049
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2050
text {*
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2051
\noindent
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2052
For recursion through curried $n$-ary functions, $n$ applications of
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2053
@{term "op \<circ>"} are necessary. The examples below illustrate the case where
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2054
$n = 2$:
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2055
*}
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2056
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2057
    codatatype ('a, 'b) sm2 =
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2058
      SM2 (accept2: bool) (trans2: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) sm2")
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2059
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2060
text {* \blankline *}
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2061
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2062
    primcorec
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2063
      (*<*)(in early) (*>*)sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) sm2"
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2064
    where
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2065
      "sm2_of_dfa \<delta> F q = SM2 (q \<in> F) (op \<circ> (op \<circ> (sm2_of_dfa \<delta> F)) (\<delta> q))"
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2066
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2067
text {* \blankline *}
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2068
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2069
    primcorec
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2070
      sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) sm2"
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2071
    where
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2072
      "sm2_of_dfa \<delta> F q = SM2 (q \<in> F) (\<lambda>a b. sm2_of_dfa \<delta> F (\<delta> q a b))"
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2073
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2074
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2075
subsubsection {* Nested-as-Mutual Corecursion
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2076
  \label{sssec:primcorec-nested-as-mutual-corecursion} *}
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2077
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2078
text {*
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2079
Just as it is possible to recurse over nested recursive datatypes as if they
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2080
were mutually recursive
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2081
(Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2082
pretend that nested codatatypes are mutually corecursive. For example:
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2083
*}
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2084
54287
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2085
(*<*)
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2086
    context late
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2087
    begin
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2088
(*>*)
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2089
    primcorec
54287
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2090
      iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2091
      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
  2092
    where
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2093
      "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" |
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2094
      "iterates\<^sub>i\<^sub>i g xs =
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2095
         (case xs of
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2096
            LNil \<Rightarrow> LNil
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2097
          | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))"
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2098
54287
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2099
text {*
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2100
\noindent
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2101
Coinduction rules are generated as
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2102
@{thm [source] iterate\<^sub>i\<^sub>i.coinduct},
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2103
@{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2104
@{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct}
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2105
and analogously for @{text strong_coinduct}. These rules and the
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2106
underlying corecursors are generated on a per-need basis and are kept in a cache
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2107
to speed up subsequent definitions.
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2108
*}
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2109
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2110
(*<*)
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2111
    end
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2112
(*>*)
7f096d8eb3d0 more docs
blanchet
parents: 54278
diff changeset
  2113
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2114
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2115
subsubsection {* Constructor View
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2116
  \label{ssec:primrec-constructor-view} *}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2117
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2118
(*<*)
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2119
    locale ctr_view
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2120
    begin
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2121
(*>*)
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2122
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2123
text {*
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2124
The constructor view is similar to the code view, but there is one separate
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2125
conditional equation per constructor rather than a single unconditional
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2126
equation. Examples that rely on a single constructor, such as @{const literate}
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2127
and @{const siterate}, are identical in both styles.
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2128
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2129
Here is an example where there is a difference:
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2130
*}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2131
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2132
    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
  2133
      "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lappend xs ys = LNil" |
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2134
      "_ \<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
  2135
         (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
  2136
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2137
text {*
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2138
\noindent
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2139
With the constructor view, we must distinguish between the @{const LNil} and
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2140
the @{const LCons} case. The condition for @{const LCons} is
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2141
left implicit, as the negation of that for @{const LNil}.
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2142
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2143
For this example, the constructor view is slighlty more involved than the
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2144
code equation. Recall the code view version presented in
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2145
Section~\ref{sssec:primcorec-simple-corecursion}.
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2146
% TODO: \[{thm code_view.lappend.code}\]
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2147
The constructor view requires us to analyze the second argument (@{term ys}).
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2148
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
  2149
% TODO: \[{thm lappend.code}\]
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2150
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2151
In contrast, the next example is arguably more naturally expressed in the
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2152
constructor view:
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2153
*}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2154
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53829
diff changeset
  2155
    primcorec
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2156
      random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2157
    where
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2158
      "n mod 4 = 0 \<Longrightarrow> random_process s f n = Fail" |
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2159
      "n mod 4 = 1 \<Longrightarrow>
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2160
         random_process s f n = Skip (random_process s f (f n))" |
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2161
      "n mod 4 = 2 \<Longrightarrow>
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2162
         random_process s f n = Action (shd s) (random_process (stl s) f (f n))" |
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2163
      "n mod 4 = 3 \<Longrightarrow>
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2164
         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
  2165
           (random_process (every_snd (stl s)) f (f n))"
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2166
(*<*)
53644
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2167
    end
eb8362530715 more (co)data docs
blanchet
parents: 53643
diff changeset
  2168
(*>*)
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2169
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2170
text {*
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2171
\noindent
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2172
Since there is no sequentiality, we can apply the equation for @{const Choice}
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2173
without having first to discharge @{term "n mod (4\<Colon>int) \<noteq> 0"},
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2174
@{term "n mod (4\<Colon>int) \<noteq> 1"}, and
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2175
@{term "n mod (4\<Colon>int) \<noteq> 2"}.
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2176
The price to pay for this elegance is that we must discharge exclusivity proof
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2177
obligations, one for each pair of conditions
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2178
@{term "(n mod (4\<Colon>int) = i, n mod (4\<Colon>int) = j)"}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2179
with @{term "i < j"}. If we prefer not to discharge any obligations, we can
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2180
enable the @{text "sequential"} option. This pushes the problem to the users of
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2181
the generated properties.
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2182
%Here are more examples to conclude:
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2183
*}
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2184
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2185
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2186
subsubsection {* Destructor View
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2187
  \label{ssec:primrec-destructor-view} *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2188
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2189
(*<*)
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2190
    locale dtr_view
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2191
    begin
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2192
(*>*)
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2193
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2194
text {*
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2195
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
  2196
determine which constructor to choose, and these conditions are interpreted
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2197
sequentially or not depending on the @{text "sequential"} option.
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2198
Consider the following examples:
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2199
*}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2200
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2201
    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
  2202
      "\<not> lnull (literate _ x)" |
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2203
      "lhd (literate _ x) = x" |
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2204
      "ltl (literate g x) = literate g (g x)"
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2205
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2206
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2207
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2208
    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
  2209
      "shd (siterate _ x) = x" |
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2210
      "stl (siterate g x) = siterate g (g x)"
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2211
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2212
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2213
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2214
    primcorec every_snd :: "'a stream \<Rightarrow> 'a stream" where
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2215
      "shd (every_snd s) = shd s" |
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2216
      "stl (every_snd s) = stl (stl s)"
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2217
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2218
text {*
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2219
\noindent
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2220
The first formula in the @{const literate} specification indicates which
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2221
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
  2222
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
  2223
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
  2224
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
  2225
Their arguments are unrestricted.
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2226
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2227
The next example shows how to specify functions that rely on more than one
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2228
constructor:
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2229
*}
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2230
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2231
    primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2232
      "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2233
      "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2234
      "ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2235
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2236
text {*
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2237
\noindent
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2238
For a codatatype with $n$ constructors, it is sufficient to specify $n - 1$
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2239
discriminator formulas. The command will then assume that the remaining
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2240
constructor should be taken otherwise. This can be made explicit by adding
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2241
*}
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2242
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2243
(*<*)
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2244
    end
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2245
54182
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2246
    locale dtr_view2
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2247
    begin
e3759cbde221 expand doc a bit
blanchet
parents: 54181
diff changeset
  2248
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2249
    primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2250
      "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2251
      "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2252
      "ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)" |
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2253
(*>*)
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2254
      "_ \<Longrightarrow> \<not> lnull (lappend xs ys)"
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2255
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2256
text {*
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2257
\noindent
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2258
to the specification. The generated selector theorems are conditional.
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2259
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2260
The next example illustrates how to cope with selectors defined for several
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2261
constructors:
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2262
*}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2263
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53829
diff changeset
  2264
    primcorec
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2265
      random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2266
    where
57091
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
  2267
      "n mod 4 = 0 \<Longrightarrow> random_process s f n = Fail" |
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2268
      "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
  2269
      "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
  2270
      "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
  2271
      "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
  2272
      "prefix (random_process s f n) = shd s" |
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53829
diff changeset
  2273
      "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
  2274
      "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
  2275
      "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
  2276
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2277
text {*
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2278
\noindent
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2279
Using the @{text "of"} keyword, different equations are specified for @{const
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2280
cont} depending on which constructor is selected.
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2281
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2282
Here are more examples to conclude:
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2283
*}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2284
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2285
    primcorec
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2286
      even_infty :: even_enat and
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2287
      odd_infty :: odd_enat
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2288
    where
57091
1fa9c19ba2c9 got rid of '=:' squiggly
blanchet
parents: 57088
diff changeset
  2289
      "even_infty \<noteq> Even_EZero" |
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2290
      "un_Even_ESuc even_infty = odd_infty" |
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2291
      "un_Odd_ESuc odd_infty = even_infty"
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2292
53752
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2293
text {* \blankline *}
1a883094fbe0 more primcorec docs
blanchet
parents: 53751
diff changeset
  2294
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2295
    primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
54072
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2296
      "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = x" |
7bee26d970f0 more primcorec docs
blanchet
parents: 54071
diff changeset
  2297
      "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = lmap (iterate\<^sub>i\<^sub>i g) (g x)"
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2298
(*<*)
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2299
    end
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2300
(*>*)
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2301
53750
03c5c2db3a47 more primcorec docs
blanchet
parents: 53749
diff changeset
  2302
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2303
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2304
  \label{ssec:primcorec-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2305
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2306
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2307
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
  2308
  \label{sssec:primcorecursive-and-primcorec} *}
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2309
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2310
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2311
\begin{matharray}{rcl}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2312
  @{command_def "primcorec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2313
  @{command_def "primcorecursive"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2314
\end{matharray}
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2315
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2316
@{rail \<open>
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 54958
diff changeset
  2317
  (@@{command primcorec} | @@{command primcorecursive}) target? \<newline>
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2318
    @{syntax pcr_option}? fixes @'where' (@{syntax pcr_formula} + '|')
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2319
  ;
53828
408c9a3617e3 improved rail diagram
blanchet
parents: 53826
diff changeset
  2320
  @{syntax_def pcr_option}: '(' ('sequential' | 'exhaustive') ')'
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2321
  ;
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2322
  @{syntax_def pcr_formula}: thmdecl? prop (@'of' (term * ))?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2323
\<close>}
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2324
55351
blanchet
parents: 55350
diff changeset
  2325
\medskip
blanchet
parents: 55350
diff changeset
  2326
blanchet
parents: 55350
diff changeset
  2327
\noindent
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2328
The @{command primcorec} and @{command primcorecursive} commands introduce a set
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2329
of mutually corecursive functions over codatatypes.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2330
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2331
The syntactic entity \synt{target} can be used to specify a local context,
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2332
\synt{fixes} denotes a list of names with optional type signatures,
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2333
\synt{thmdecl} denotes an optional name for the formula that follows, and
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2334
\synt{prop} denotes a HOL proposition \cite{isabelle-isar-ref}.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2335
56124
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
  2336
The optional target is optionally followed by one or both of the following
315cc3c0a19a tuned wording (pun)
blanchet
parents: 56123
diff changeset
  2337
options:
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2338
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2339
\begin{itemize}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2340
\setlength{\itemsep}{0pt}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2341
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2342
\item
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2343
The @{text "sequential"} option indicates that the conditions in specifications
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2344
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
  2345
sequentially.
53826
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2346
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2347
\item
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2348
The @{text "exhaustive"} option indicates that the conditions in specifications
92a8ae172242 use "primcorec" in doc
blanchet
parents: 53822
diff changeset
  2349
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
  2350
\end{itemize}
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2351
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2352
The @{command primcorec} command is an abbreviation for @{command
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2353
primcorecursive} with @{text "by auto?"} to discharge any emerging proof
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2354
obligations.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2355
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2356
%%% TODO: elaborate on format of the propositions
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2357
%%% TODO: elaborate on mutual and nested-to-mutual
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2358
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2359
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2360
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2361
(*
52840
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2362
subsection {* Generated Theorems
a0da63cec918 more (co)datatype documentation
blanchet
parents: 52829
diff changeset
  2363
  \label{ssec:primcorec-generated-theorems} *}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2364
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2365
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2366
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2367
(*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2368
subsection {* Recursive Default Values for Selectors
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2369
  \label{ssec:primcorec-recursive-default-values-for-selectors} *}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2370
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2371
text {*
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2372
partial_function to the rescue
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2373
*}
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2374
*)
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2375
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2376
55351
blanchet
parents: 55350
diff changeset
  2377
section {* Introducing Bounded Natural Functors
blanchet
parents: 55350
diff changeset
  2378
  \label{sec:introducing-bounded-natural-functors} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  2379
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2380
text {*
53647
e78ebb290dd6 more (co)data docs
blanchet
parents: 53646
diff changeset
  2381
The (co)datatype package can be set up to allow nested recursion through
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2382
arbitrary type constructors, as long as they adhere to the BNF requirements
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2383
and are registered as BNFs. It is also possible to declare a BNF abstractly
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2384
without specifying its internal structure.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2385
*}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2386
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2387
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2388
subsection {* Bounded Natural Functors
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2389
  \label{ssec:bounded-natural-functors} *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2390
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2391
text {*
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2392
Bounded natural functors (BNFs) are a semantic criterion for where
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2393
(co)re\-cur\-sion may appear on the right-hand side of an equation
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2394
\cite{traytel-et-al-2012,blanchette-et-al-wit}.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2395
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2396
An $n$-ary BNF is a type constructor equipped with a map function
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2397
(functorial action), $n$ set functions (natural transformations),
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2398
and an infinite cardinal bound that satisfy certain properties.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2399
For example, @{typ "'a llist"} is a unary BNF.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2400
Its relator @{text "llist_all2 \<Colon>
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2401
  ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow>
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2402
  'a llist \<Rightarrow> 'b llist \<Rightarrow> bool"}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2403
extends binary predicates over elements to binary predicates over parallel
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2404
lazy lists. The cardinal bound limits the number of elements returned by the
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2405
set function; it may not depend on the cardinality of @{typ 'a}.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2406
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2407
The type constructors introduced by @{command datatype_new} and
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2408
@{command codatatype} are automatically registered as BNFs. In addition, a
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2409
number of old-style datatypes and non-free types are preregistered.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2410
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2411
Given an $n$-ary BNF, the $n$ type variables associated with set functions,
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2412
and on which the map function acts, are \emph{live}; any other variables are
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2413
\emph{dead}. Nested (co)recursion can only take place through live variables.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2414
*}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2415
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2416
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2417
subsection {* Introductory Examples
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2418
  \label{ssec:bnf-introductory-examples} *}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2419
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2420
text {*
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2421
The example below shows how to register a type as a BNF using the @{command bnf}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2422
command. Some of the proof obligations are best viewed with the theory
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2423
@{theory Cardinal_Notations}, located in \verb|~~/src/HOL/Library|,
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2424
imported.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2425
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2426
The type is simply a copy of the function space @{typ "'d \<Rightarrow> 'a"}, where @{typ 'a}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2427
is live and @{typ 'd} is dead. We introduce it together with its map function,
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2428
set function, and relator.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2429
*}
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2430
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2431
    typedef ('d, 'a) fn = "UNIV \<Colon> ('d \<Rightarrow> 'a) set"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2432
    by simp
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2433
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2434
    text {* \blankline *}
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2435
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2436
    setup_lifting type_definition_fn
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2437
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2438
text {* \blankline *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2439
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2440
    lift_definition map_fn :: "('a \<Rightarrow> 'b) \<Rightarrow> ('d, 'a) fn \<Rightarrow> ('d, 'b) fn" is "op \<circ>" .
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2441
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2442
text {* \blankline *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2443
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2444
    lift_definition set_fn :: "('d, 'a) fn \<Rightarrow> 'a set" is range .
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2445
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2446
text {* \blankline *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2447
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2448
    lift_definition
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2449
      rel_fn :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('d, 'a) fn \<Rightarrow> ('d, 'b) fn \<Rightarrow> bool"
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2450
    is
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55896
diff changeset
  2451
      "rel_fun (op =)" .
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2452
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2453
text {* \blankline *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2454
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2455
    bnf "('d, 'a) fn"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2456
      map: map_fn
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2457
      sets: set_fn
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2458
      bd: "natLeq +c |UNIV :: 'd set|"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2459
      rel: rel_fn
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2460
    proof -
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2461
      show "map_fn id = id"
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2462
        by transfer auto
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2463
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2464
      fix F G show "map_fn (G \<circ> F) = map_fn G \<circ> map_fn F"
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2465
        by transfer (auto simp add: comp_def)
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2466
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2467
      fix F f g
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2468
      assume "\<And>x. x \<in> set_fn F \<Longrightarrow> f x = g x"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2469
      thus "map_fn f F = map_fn g F"
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2470
        by transfer auto
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2471
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2472
      fix f show "set_fn \<circ> map_fn f = op ` f \<circ> set_fn"
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2473
        by transfer (auto simp add: comp_def)
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2474
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2475
      show "card_order (natLeq +c |UNIV \<Colon> 'd set| )"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2476
        apply (rule card_order_csum)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2477
        apply (rule natLeq_card_order)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2478
        by (rule card_of_card_order_on)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2479
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2480
      show "cinfinite (natLeq +c |UNIV \<Colon> 'd set| )"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2481
        apply (rule cinfinite_csum)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2482
        apply (rule disjI1)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2483
        by (rule natLeq_cinfinite)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2484
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2485
      fix F :: "('d, 'a) fn"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2486
      have "|set_fn F| \<le>o |UNIV \<Colon> 'd set|" (is "_ \<le>o ?U")
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2487
        by transfer (rule card_of_image)
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2488
      also have "?U \<le>o natLeq +c ?U"
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2489
        by (rule ordLeq_csum2) (rule card_of_Card_order)
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2490
      finally show "|set_fn F| \<le>o natLeq +c |UNIV \<Colon> 'd set|" .
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2491
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2492
      fix R S
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2493
      show "rel_fn R OO rel_fn S \<le> rel_fn (R OO S)"
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55896
diff changeset
  2494
        by (rule, transfer) (auto simp add: rel_fun_def)
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2495
    next
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2496
      fix R
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2497
      show "rel_fn R =
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2498
            (BNF_Util.Grp {x. set_fn x \<subseteq> Collect (split R)} (map_fn fst))\<inverse>\<inverse> OO
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2499
             BNF_Util.Grp {x. set_fn x \<subseteq> Collect (split R)} (map_fn snd)"
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2500
        unfolding Grp_def fun_eq_iff relcompp.simps conversep.simps
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2501
        apply transfer
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55896
diff changeset
  2502
        unfolding rel_fun_def subset_iff image_iff
55459
1cd927ca8296 cleaner, complete proof in documentation, contributed by Dmitriy T.
blanchet
parents: 55410
diff changeset
  2503
        by auto (force, metis pair_collapse)
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2504
    qed
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2505
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2506
text {* \blankline *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2507
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2508
    print_theorems
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2509
    print_bnfs
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2510
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2511
text {*
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2512
\noindent
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2513
Using \keyw{print\_theorems} and @{keyword print_bnfs}, we can contemplate and
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2514
show the world what we have achieved.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2515
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2516
This particular example does not need any nonemptiness witness, because the
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2517
one generated by default is good enough, but in general this would be
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2518
necessary. See \verb|~~/src/HOL/Basic_BNFs.thy|,
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2519
\verb|~~/src/HOL/Library/FSet.thy|, and \verb|~~/src/HOL/Library/Multiset.thy|
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2520
for further examples of BNF registration, some of which feature custom
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2521
witnesses.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2522
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
  2523
The next example declares a BNF axiomatically. The @{command bnf_axiomatization} command
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2524
introduces a type @{text "('a, 'b, 'c) F"}, three set constants, a map
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2525
function, a relator, and a nonemptiness witness that depends only on
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2526
@{typ 'a}. (The type @{text "'a \<Rightarrow> ('a, 'b, 'c) F"} of
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2527
the witness can be read as an implication: If we have a witness for @{typ 'a},
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2528
we can construct a witness for @{text "('a, 'b, 'c) F"}.) The BNF
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2529
properties are postulated as axioms.
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2530
*}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2531
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
  2532
    bnf_axiomatization (setA: 'a, setB: 'b, setC: 'c) F [wits: "'a \<Rightarrow> ('a, 'b, 'c) F"]
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2533
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2534
text {* \blankline *}
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2535
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2536
    print_theorems
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2537
    print_bnfs
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2538
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2539
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2540
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2541
  \label{ssec:bnf-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2542
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2543
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  2544
subsubsection {* \keyw{bnf}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  2545
  \label{sssec:bnf} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2546
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  2547
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2548
\begin{matharray}{rcl}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2549
  @{command_def "bnf"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2550
\end{matharray}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2551
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2552
@{rail \<open>
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2553
  @@{command bnf} target? (name ':')? type \<newline>
55029
61a6bf7d4b02 clarified @{rail} syntax: prefer explicit \<newline> symbol;
wenzelm
parents: 54958
diff changeset
  2554
    'map:' term ('sets:' (term +))? 'bd:' term \<newline>
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54402
diff changeset
  2555
    ('wits:' (term +))? ('rel:' term)?
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2556
\<close>}
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2557
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2558
\medskip
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2559
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2560
\noindent
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2561
The @{command bnf} command registers an existing type as a bounded natural
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2562
functor (BNF). The type must be equipped with an appropriate map function
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2563
(functorial action). In addition, custom set functions, relators, and
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2564
nonemptiness witnesses can be specified; otherwise, default versions are used.
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2565
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2566
The syntactic entity \synt{target} can be used to specify a local context,
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2567
\synt{type} denotes a HOL type, and \synt{term} denotes a HOL term
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2568
\cite{isabelle-isar-ref}.
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2569
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2570
%%% TODO: elaborate on proof obligations
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  2571
*}
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2572
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2573
56948
1144d7ec892a more bnf_decl -> bnf_axiomatization
blanchet
parents: 56947
diff changeset
  2574
subsubsection {* \keyw{bnf\_axiomatization}
1144d7ec892a more bnf_decl -> bnf_axiomatization
blanchet
parents: 56947
diff changeset
  2575
  \label{sssec:bnf-axiomatization} *}
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2576
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2577
text {*
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2578
\begin{matharray}{rcl}
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
  2579
  @{command_def "bnf_axiomatization"} & : & @{text "local_theory \<rightarrow> local_theory"}
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2580
\end{matharray}
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2581
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2582
@{rail \<open>
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  2583
  @@{command bnf_axiomatization} target? @{syntax tyargs}? name @{syntax wit_types}? \<newline>
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  2584
    mixfix? @{syntax map_rel}?
54602
168790252038 some documentation
traytel
parents: 54537
diff changeset
  2585
  ;
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2586
  @{syntax_def wit_types}: '[' 'wits' ':' types ']'
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2587
\<close>}
54602
168790252038 some documentation
traytel
parents: 54537
diff changeset
  2588
55351
blanchet
parents: 55350
diff changeset
  2589
\medskip
blanchet
parents: 55350
diff changeset
  2590
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2591
\noindent
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
  2592
The @{command bnf_axiomatization} command declares a new type and associated constants
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2593
(map, set, relator, and cardinal bound) and asserts the BNF properties for
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2594
these constants as axioms.
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2595
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2596
The syntactic entity \synt{target} can be used to specify a local context,
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2597
\synt{name} denotes an identifier, \synt{typefree} denotes fixed type variable
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2598
(@{typ 'a}, @{typ 'b}, \ldots), and \synt{mixfix} denotes the usual
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2599
parenthesized mixfix notation \cite{isabelle-isar-ref}.
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2600
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2601
Type arguments are live by default; they can be marked as dead by entering
57092
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
  2602
``@{text dead}'' in front of the type variable (e.g., ``@{text "(dead 'a)"}'')
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
  2603
instead of an identifier for the corresponding set function. Witnesses can be
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
  2604
specified by their types. Otherwise, the syntax of @{command bnf_axiomatization}
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
  2605
is identical to the left-hand side of a @{command datatype_new} or
59603f4f1d2e changed '-:' to 'dead' in BNF
blanchet
parents: 57091
diff changeset
  2606
@{command codatatype} definition.
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2607
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2608
The command is useful to reason abstractly about BNFs. The axioms are safe
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2609
because there exists BNFs of arbitrary large arities. Applications must import
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 56904
diff changeset
  2610
the theory @{theory BNF_Axiomatization}, located in the directory
55350
cf34abe28209 docs about registering a BNF
blanchet
parents: 55290
diff changeset
  2611
\verb|~~/src/HOL/Library|, to use this functionality.
54187
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2612
*}
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2613
80259d79a81e more doc -- feedback from Andrei P.
blanchet
parents: 54185
diff changeset
  2614
53621
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  2615
subsubsection {* \keyw{print\_bnfs}
9c3a80af72ff more (co)data doc
blanchet
parents: 53619
diff changeset
  2616
  \label{sssec:print-bnfs} *}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2618
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2619
\begin{matharray}{rcl}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2620
  @{command_def "print_bnfs"} & : & @{text "local_theory \<rightarrow>"}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2621
\end{matharray}
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2622
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2623
@{rail \<open>
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2624
  @@{command print_bnfs}
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2625
\<close>}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2626
*}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2627
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2628
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2629
section {* Deriving Destructors and Theorems for Free Constructors
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2630
  \label{sec:deriving-destructors-and-theorems-for-free-constructors} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2631
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2632
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2633
The derivation of convenience theorems for types equipped with free constructors,
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2634
as performed internally by @{command datatype_new} and @{command codatatype},
55468
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
  2635
is available as a stand-alone command called @{command free_constructors}.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2636
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2637
%  * 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
  2638
%    a type not introduced by ...
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2639
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2640
%  * also useful for compatibility with old package, e.g. add destructors to
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2641
%    old \keyw{datatype}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2642
%
55468
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
  2643
%  * @{command free_constructors}
55410
54b09e82b9e1 killed 'rep_compat' option
blanchet
parents: 55398
diff changeset
  2644
%    * @{text "no_discs_sels"}, @{text "no_code"}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2645
%    * hack to have both co and nonco view via locale (cf. ext nats)
54616
a21a2223c02b minor doc update
blanchet
parents: 54602
diff changeset
  2646
%  * code generator
a21a2223c02b minor doc update
blanchet
parents: 54602
diff changeset
  2647
%     * eq, refl, simps
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2648
*}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  2649
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2650
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2651
(*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2652
subsection {* Introductory Example
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2653
  \label{ssec:ctors-introductory-example} *}
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2654
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2655
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2656
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2657
subsection {* Command Syntax
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2658
  \label{ssec:ctors-command-syntax} *}
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2659
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2660
55472
990651bfc65b updated docs to reflect the new 'free_constructors' syntax
blanchet
parents: 55468
diff changeset
  2661
subsubsection {* \keyw{free\_constructors}
990651bfc65b updated docs to reflect the new 'free_constructors' syntax
blanchet
parents: 55468
diff changeset
  2662
  \label{sssec:free-constructors} *}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2663
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  2664
text {*
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2665
\begin{matharray}{rcl}
55468
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
  2666
  @{command_def "free_constructors"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
53829
92e71eb22ebe more (co)data docs
blanchet
parents: 53828
diff changeset
  2667
\end{matharray}
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  2668
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2669
@{rail \<open>
55468
98b25c51e9e5 renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents: 55460
diff changeset
  2670
  @@{command free_constructors} target? @{syntax dt_options} \<newline>
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  2671
    name 'for' (@{syntax fc_ctor} + '|') \<newline>
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 57200
diff changeset
  2672
  (@'where' (prop + '|'))?
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  2673
  ;
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  2674
  @{syntax_def fc_ctor}: (name ':')? term (name * )
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 55029
diff changeset
  2675
\<close>}
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  2676
55351
blanchet
parents: 55350
diff changeset
  2677
\medskip
blanchet
parents: 55350
diff changeset
  2678
55460
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2679
\noindent
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2680
The @{command free_constructors} command generates destructor constants for
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2681
freely constructed types as well as properties about constructors and
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2682
destructors. It also registers the constants and theorems in a data structure
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2683
that is queried by various tools (e.g., \keyw{function}).
3f4efd7d950d added a bit of documentation for the different commands
blanchet
parents: 55459
diff changeset
  2684
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2685
The syntactic entity \synt{target} can be used to specify a local context,
57200
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  2686
\synt{name} denotes an identifier, \synt{prop} denotes a HOL proposition, and
aab87ffa60cc use 'where' clause for selector default value syntax
blanchet
parents: 57153
diff changeset
  2687
\synt{term} denotes a HOL term \cite{isabelle-isar-ref}.
55474
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2688
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2689
The syntax resembles that of @{command datatype_new} and @{command codatatype}
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2690
definitions (Sections
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2691
\ref{ssec:datatype-command-syntax}~and~\ref{ssec:codatatype-command-syntax}).
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2692
A constructor is specified by an optional name for the discriminator, the
501df9ad117b more (co)datatype docs
blanchet
parents: 55472
diff changeset
  2693
constructor itself (as a term), and a list of optional names for the selectors.
53028
a1e64c804c35 more (co)datatype documentation
blanchet
parents: 53025
diff changeset
  2694
53542
14000a283ce0 more (co)data docs
blanchet
parents: 53535
diff changeset
  2695
Section~\ref{ssec:datatype-generated-theorems} lists the generated theorems.
55398
67e9fdd9ae9e updated docs
blanchet
parents: 55355
diff changeset
  2696
For technical reasons, the @{text "[fundef_cong]"} attribute is not set on the
67e9fdd9ae9e updated docs
blanchet
parents: 55355
diff changeset
  2697
generated @{text case_cong} theorem. It can be added manually using
67e9fdd9ae9e updated docs
blanchet
parents: 55355
diff changeset
  2698
\keyw{declare}.
53018
11ebef554439 added rail diagram
blanchet
parents: 52969
diff changeset
  2699
*}
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2700
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2701
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2702
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  2703
section {* Standard ML Interface
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2704
  \label{sec:standard-ml-interface} *}
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  2705
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2706
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2707
The package's programmatic interface.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2708
*}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2709
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2710
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2711
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2712
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  2713
section {* Interoperability
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2714
  \label{sec:interoperability} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2715
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2716
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2717
The package's interaction with other Isabelle packages and tools, such as the
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2718
code generator and the counterexample generators.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2719
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2720
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2721
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2722
subsection {* Transfer and Lifting
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2723
  \label{ssec:transfer-and-lifting} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2724
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2725
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2726
subsection {* Code Generator
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2727
  \label{ssec:code-generator} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2728
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2729
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2730
subsection {* Quickcheck
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2731
  \label{ssec:quickcheck} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2732
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2733
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2734
subsection {* Nitpick
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2735
  \label{ssec:nitpick} *}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2736
52824
b7a83845bc93 more (co)datatype documentation
blanchet
parents: 52822
diff changeset
  2737
52828
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2738
subsection {* Nominal Isabelle
e1c6fa322d96 more (co)datatype docs
blanchet
parents: 52827
diff changeset
  2739
  \label{ssec:nominal-isabelle} *}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2740
*)
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2741
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2742
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2743
(*
52827
395d3df496ed more (co)datatype documentation
blanchet
parents: 52824
diff changeset
  2744
section {* Known Bugs and Limitations
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2745
  \label{sec:known-bugs-and-limitations} *}
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2746
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2747
text {*
53623
501e2091182b more (co)data docs
blanchet
parents: 53621
diff changeset
  2748
Known open issues of the package.
52805
7f2f42046361 more (co)datatype documentation
blanchet
parents: 52795
diff changeset
  2749
*}
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2750
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2751
text {*
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2752
%* 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
  2753
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2754
%* partial documentation
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2755
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2756
%* no way to register "sum" and "prod" as (co)datatypes to enable N2M reduction for them
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
  2757
%  (for @{command datatype_compat} and prim(co)rec)
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2758
%
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2759
%    * a fortiori, no way to register same type as both data- and codatatype
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2760
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2761
%* no recursion through unused arguments (unlike with the old package)
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2762
%
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2763
%* in a locale, cannot use locally fixed types (because of limitation in typedef)?
53619
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2764
%
27d2c98d9d9f more (co)data docs
blanchet
parents: 53617
diff changeset
  2765
% *names of variables suboptimal
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  2766
*}
53675
d4a4b32038eb more (co)data docs
blanchet
parents: 53654
diff changeset
  2767
*)
52822
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  2768
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  2769
ae938ac9a721 more (co)datatype docs
blanchet
parents: 52806
diff changeset
  2770
text {*
53863
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
  2771
\section*{Acknowledgment}
c7364dca96f2 textual improvements following Christian Sternagel's feedback
blanchet
parents: 53857
diff changeset
  2772
53749
b37db925b663 adapted primcorec documentation to reflect the three views
blanchet
parents: 53748
diff changeset
  2773
Tobias Nipkow and Makarius Wenzel encouraged us to implement the new
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2774
(co)datatype package. Andreas Lochbihler provided lots of comments on earlier
56655
34f2fe40dc7b updated BNF docs
blanchet
parents: 56653
diff changeset
  2775
versions of the package, especially on the coinductive part. Brian Huffman
55129
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
  2776
suggested major simplifications to the internal constructions, many of which
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
  2777
have yet to be implemented. Florian Haftmann and Christian Urban provided
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
  2778
general advice on Isabelle and package writing. Stefan Milius and Lutz
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
  2779
Schr\"oder found an elegant proof that eliminated one of the BNF proof
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
  2780
obligations. Andreas Lochbihler and Christian Sternagel suggested many textual
26bd1cba3ab5 killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents: 55114
diff changeset
  2781
improvements to this tutorial.
52794
aae782070611 more (co)datatype documentation
blanchet
parents: 52792
diff changeset
  2782
*}
53617
da5e1887d7a7 more (co)data doc
blanchet
parents: 53565
diff changeset
  2783
52792
3e651be14fcd sketched documentation for new (co)datatype package
blanchet
parents:
diff changeset
  2784
end