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