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