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