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