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