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