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