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