diff -r bda7527ccf05 -r cc2d676d5395 src/Doc/Datatypes/Datatypes.thy --- a/src/Doc/Datatypes/Datatypes.thy Wed Dec 26 16:07:28 2018 +0100 +++ b/src/Doc/Datatypes/Datatypes.thy Wed Dec 26 16:25:20 2018 +0100 @@ -85,7 +85,7 @@ (co)datatypes are derived rather than introduced axiomatically.% \footnote{However, some of the internal constructions and most of the internal proof obligations are omitted -if the @{text quick_and_dirty} option is enabled.} +if the \quick_and_dirty\ option is enabled.} The package is described in a number of scientific papers @{cite "traytel-et-al-2012" and "blanchette-et-al-2014-impl" and "panny-et-al-2014" and "blanchette-et-al-2015-wit"}. @@ -194,8 +194,8 @@ text \ \noindent -The constructors are @{text "None :: 'a option"} and -@{text "Some :: 'a \ 'a option"}. +The constructors are \None :: 'a option\ and +\Some :: 'a \ 'a option\. The next example has three type parameters: \ @@ -205,7 +205,7 @@ text \ \noindent The constructor is -@{text "Triple :: 'a \ 'b \ 'c \ ('a, 'b, 'c) triple"}. +\Triple :: 'a \ 'b \ 'c \ ('a, 'b, 'c) triple\. Unlike in Standard ML, curried constructors are supported. The uncurried variant is also possible: \ @@ -286,9 +286,9 @@ text \ \noindent -The issue is that the function arrow @{text "\"} allows recursion +The issue is that the function arrow \\\ allows recursion only through its right-hand side. This issue is inherited by polymorphic -datatypes defined in terms of~@{text "\"}: +datatypes defined in terms of~\\\: \ datatype ('a, 'b) fun_copy = Fun "'a \ 'b" @@ -311,9 +311,9 @@ text \ \noindent -In general, type constructors @{text "('a\<^sub>1, \, 'a\<^sub>m) t"} -allow recursion on a subset of their type arguments @{text 'a\<^sub>1}, \ldots, -@{text 'a\<^sub>m}. These type arguments are called \emph{live}; the remaining +In general, type constructors \('a\<^sub>1, \, 'a\<^sub>m) t\ +allow recursion on a subset of their type arguments \'a\<^sub>1\, \ldots, +\'a\<^sub>m\. These type arguments are called \emph{live}; the remaining type arguments are called \emph{dead}. In @{typ "'a \ 'b"} and @{typ "('a, 'b) fun_copy"}, the type variable @{typ 'a} is dead and @{typ 'b} is live. @@ -345,10 +345,10 @@ The @{command datatype} command introduces various constants in addition to the constructors. With each datatype are associated set functions, a map function, a predicator, a relator, discriminators, and selectors, all of which can be given -custom names. In the example below, the familiar names @{text null}, @{text hd}, -@{text tl}, @{text set}, @{text map}, and @{text list_all2} override the -default names @{text is_Nil}, @{text un_Cons1}, @{text un_Cons2}, -@{text set_list}, @{text map_list}, and @{text rel_list}: +custom names. In the example below, the familiar names \null\, \hd\, +\tl\, \set\, \map\, and \list_all2\ override the +default names \is_Nil\, \un_Cons1\, \un_Cons2\, +\set_list\, \map_list\, and \rel_list\: \ (*<*) @@ -384,21 +384,21 @@ \begin{tabular}{@ {}ll@ {}} Constructors: & - @{text "Nil :: 'a list"} \\ + \Nil :: 'a list\ \\ & - @{text "Cons :: 'a \ 'a list \ 'a list"} \\ + \Cons :: 'a \ 'a list \ 'a list\ \\ Discriminator: & - @{text "null :: 'a list \ bool"} \\ + \null :: 'a list \ bool\ \\ Selectors: & - @{text "hd :: 'a list \ 'a"} \\ + \hd :: 'a list \ 'a\ \\ & - @{text "tl :: 'a list \ 'a list"} \\ + \tl :: 'a list \ 'a list\ \\ Set function: & - @{text "set :: 'a list \ 'a set"} \\ + \set :: 'a list \ 'a set\ \\ Map function: & - @{text "map :: ('a \ 'b) \ 'a list \ 'b list"} \\ + \map :: ('a \ 'b) \ 'a list \ 'b list\ \\ Relator: & - @{text "list_all2 :: ('a \ 'b \ bool) \ 'a list \ 'b list \ bool"} + \list_all2 :: ('a \ 'b \ bool) \ 'a list \ 'b list \ bool\ \end{tabular} \medskip @@ -466,7 +466,7 @@ text \ \begin{matharray}{rcl} - @{command_def "datatype"} & : & @{text "local_theory \ local_theory"} + @{command_def "datatype"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -489,7 +489,7 @@ datatypes specified by their constructors. The syntactic entity \synt{target} can be used to specify a local -context (e.g., @{text "(in linorder)"} @{cite "isabelle-isar-ref"}), +context (e.g., \(in linorder)\ @{cite "isabelle-isar-ref"}), and \synt{prop} denotes a HOL proposition. The optional target is optionally followed by a combination of the following @@ -499,19 +499,19 @@ \setlength{\itemsep}{0pt} \item -The @{text plugins} option indicates which plugins should be enabled -(@{text only}) or disabled (@{text del}). By default, all plugins are enabled. +The \plugins\ option indicates which plugins should be enabled +(\only\) or disabled (\del\). By default, all plugins are enabled. \item -The @{text "discs_sels"} option indicates that discriminators and selectors +The \discs_sels\ option indicates that discriminators and selectors should be generated. The option is implicitly enabled if names are specified for discriminators or selectors. \end{itemize} The optional \keyw{where} clause specifies default values for selectors. Each proposition must be an equation of the form -@{text "un_D (C \) = \"}, where @{text C} is a constructor and -@{text un_D} is a selector. +\un_D (C \) = \\, where \C\ is a constructor and +\un_D\ is a selector. The left-hand sides of the datatype equations specify the name of the type to define, its type parameters, and additional information: @@ -530,9 +530,9 @@ variable (@{typ 'a}, @{typ 'b}, \ldots) @{cite "isabelle-isar-ref"}. The optional names preceding the type variables allow to override the default -names of the set functions (@{text set\<^sub>1_t}, \ldots, @{text set\<^sub>m_t}). Type -arguments can be marked as dead by entering @{text dead} in front of the -type variable (e.g., @{text "(dead 'a)"}); otherwise, they are live or dead +names of the set functions (\set\<^sub>1_t\, \ldots, \set\<^sub>m_t\). Type +arguments can be marked as dead by entering \dead\ in front of the +type variable (e.g., \(dead 'a)\); otherwise, they are live or dead (and a set function is generated or not) depending on where they occur in the right-hand sides of the definition. Declaring a type argument as dead can speed up the type definition but will prevent any later (co)recursion through that @@ -551,9 +551,9 @@ The main constituents of a constructor specification are the name of the constructor and the list of its argument types. An optional discriminator name can be supplied at the front. If discriminators are enabled (cf.\ the -@{text "discs_sels"} option) but no name is supplied, the default is -@{text "\x. x = C\<^sub>j"} for nullary constructors and -@{text t.is_C\<^sub>j} otherwise. +\discs_sels\ option) but no name is supplied, the default is +\\x. x = C\<^sub>j\ for nullary constructors and +\t.is_C\<^sub>j\ otherwise. @{rail \ @{syntax_def dt_ctor_arg}: type | '(' name ':' type ')' @@ -567,8 +567,8 @@ In addition to the type of a constructor argument, it is possible to specify a name for the corresponding selector. The same selector name can be reused for arguments to several constructors as long as the arguments share the same type. -If selectors are enabled (cf.\ the @{text "discs_sels"} option) but no name is -supplied, the default name is @{text un_C\<^sub>ji}. +If selectors are enabled (cf.\ the \discs_sels\ option) but no name is +supplied, the default name is \un_C\<^sub>ji\. \ @@ -577,7 +577,7 @@ text \ \begin{matharray}{rcl} - @{command_def "datatype_compat"} & : & @{text "local_theory \ local_theory"} + @{command_def "datatype_compat"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -609,12 +609,12 @@ \setlength{\itemsep}{0pt} \item The old-style, nested-as-mutual induction rule and recursor theorems are -generated under their usual names but with ``@{text "compat_"}'' prefixed -(e.g., @{text compat_tree.induct}, @{text compat_tree.inducts}, and -@{text compat_tree.rec}). These theorems should be identical to the ones +generated under their usual names but with ``\compat_\'' prefixed +(e.g., \compat_tree.induct\, \compat_tree.inducts\, and +\compat_tree.rec\). These theorems should be identical to the ones generated by the old datatype package, \emph{up to the order of the -premises}---meaning that the subgoals generated by the @{text induct} or -@{text induction} method may be in a different order than before. +premises}---meaning that the subgoals generated by the \induct\ or +\induction\ method may be in a different order than before. \item All types through which recursion takes place must be new-style datatypes or the function type. @@ -626,42 +626,42 @@ \label{ssec:datatype-generated-constants}\ text \ -Given a datatype @{text "('a\<^sub>1, \, 'a\<^sub>m) t"} with $m$ live type variables -and $n$ constructors @{text "t.C\<^sub>1"}, \ldots, @{text "t.C\<^sub>n"}, the following +Given a datatype \('a\<^sub>1, \, 'a\<^sub>m) t\ with $m$ live type variables +and $n$ constructors \t.C\<^sub>1\, \ldots, \t.C\<^sub>n\, the following auxiliary constants are introduced: \medskip \begin{tabular}{@ {}ll@ {}} Case combinator: & - @{text t.case_t} (rendered using the familiar @{text case}--@{text of} syntax) \\ + \t.case_t\ (rendered using the familiar \case\--\of\ syntax) \\ Discriminators: & - @{text t.is_C\<^sub>1}$, \ldots, $@{text t.is_C\<^sub>n} \\ + \t.is_C\<^sub>1\$, \ldots, $\t.is_C\<^sub>n\ \\ Selectors: & - @{text t.un_C\<^sub>11}$, \ldots, $@{text t.un_C\<^sub>1k\<^sub>1} \\ + \t.un_C\<^sub>11\$, \ldots, $\t.un_C\<^sub>1k\<^sub>1\ \\ & \quad\vdots \\ -& @{text t.un_C\<^sub>n1}$, \ldots, $@{text t.un_C\<^sub>nk\<^sub>n} \\ +& \t.un_C\<^sub>n1\$, \ldots, $\t.un_C\<^sub>nk\<^sub>n\ \\ Set functions: & - @{text t.set\<^sub>1_t}, \ldots, @{text t.set\<^sub>m_t} \\ + \t.set\<^sub>1_t\, \ldots, \t.set\<^sub>m_t\ \\ Map function: & - @{text t.map_t} \\ + \t.map_t\ \\ Relator: & - @{text t.rel_t} \\ + \t.rel_t\ \\ Recursor: & - @{text t.rec_t} + \t.rec_t\ \end{tabular} \medskip \noindent -The discriminators and selectors are generated only if the @{text "discs_sels"} +The discriminators and selectors are generated only if the \discs_sels\ option is enabled or if names are specified for discriminators or selectors. The set functions, map function, predicator, and relator are generated only if $m > 0$. In addition, some of the plugins introduce their own constants (Section~\ref{sec:selecting-plugins}). The case combinator, discriminators, and selectors are collectively called \emph{destructors}. The prefix -``@{text "t."}'' is an optional component of the names and is normally hidden. +``\t.\'' is an optional component of the names and is normally hidden. \ @@ -717,17 +717,17 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{inject} @{text "[iff, induct_simp]"}\rm:] ~ \\ +\item[\t.\\hthm{inject} \[iff, induct_simp]\\rm:] ~ \\ @{thm list.inject[no_vars]} -\item[@{text "t."}\hthm{distinct} @{text "[simp, induct_simp]"}\rm:] ~ \\ +\item[\t.\\hthm{distinct} \[simp, induct_simp]\\rm:] ~ \\ @{thm list.distinct(1)[no_vars]} \\ @{thm list.distinct(2)[no_vars]} -\item[@{text "t."}\hthm{exhaust} @{text "[cases t, case_names C\<^sub>1 \ C\<^sub>n]"}\rm:] ~ \\ +\item[\t.\\hthm{exhaust} \[cases t, case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust[no_vars]} -\item[@{text "t."}\hthm{nchotomy}\rm:] ~ \\ +\item[\t.\\hthm{nchotomy}\rm:] ~ \\ @{thm list.nchotomy[no_vars]} \end{description} @@ -739,7 +739,7 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{distinct {\upshape[}THEN notE}@{text ", elim!"}\hthm{\upshape]}\rm:] ~ \\ +\item[\t.\\hthm{distinct {\upshape[}THEN notE}\, elim!\\hthm{\upshape]}\rm:] ~ \\ @{thm list.distinct(1)[THEN notE, elim!, no_vars]} \\ @{thm list.distinct(2)[THEN notE, elim!, no_vars]} @@ -752,28 +752,28 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{case} @{text "[simp, code]"}\rm:] ~ \\ +\item[\t.\\hthm{case} \[simp, code]\\rm:] ~ \\ @{thm list.case(1)[no_vars]} \\ @{thm list.case(2)[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "t."}\hthm{case_cong} @{text "[fundef_cong]"}\rm:] ~ \\ +\item[\t.\\hthm{case_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.case_cong[no_vars]} -\item[@{text "t."}\hthm{case_cong_weak} @{text "[cong]"}\rm:] ~ \\ +\item[\t.\\hthm{case_cong_weak} \[cong]\\rm:] ~ \\ @{thm list.case_cong_weak[no_vars]} -\item[@{text "t."}\hthm{case_distrib}\rm:] ~ \\ +\item[\t.\\hthm{case_distrib}\rm:] ~ \\ @{thm list.case_distrib[no_vars]} -\item[@{text "t."}\hthm{split}\rm:] ~ \\ +\item[\t.\\hthm{split}\rm:] ~ \\ @{thm list.split[no_vars]} -\item[@{text "t."}\hthm{split_asm}\rm:] ~ \\ +\item[\t.\\hthm{split_asm}\rm:] ~ \\ @{thm list.split_asm[no_vars]} -\item[@{text "t."}\hthm{splits} = @{text "split split_asm"}] +\item[\t.\\hthm{splits} = \split split_asm\] \end{description} \end{indentblock} @@ -784,55 +784,55 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{disc} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{disc} \[simp]\\rm:] ~ \\ @{thm list.disc(1)[no_vars]} \\ @{thm list.disc(2)[no_vars]} -\item[@{text "t."}\hthm{discI}\rm:] ~ \\ +\item[\t.\\hthm{discI}\rm:] ~ \\ @{thm list.discI(1)[no_vars]} \\ @{thm list.discI(2)[no_vars]} -\item[@{text "t."}\hthm{sel} @{text "[simp, code]"}\rm:] ~ \\ +\item[\t.\\hthm{sel} \[simp, code]\\rm:] ~ \\ @{thm list.sel(1)[no_vars]} \\ @{thm list.sel(2)[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "t."}\hthm{collapse} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{collapse} \[simp]\\rm:] ~ \\ @{thm list.collapse(1)[no_vars]} \\ @{thm list.collapse(2)[no_vars]} \\ -The @{text "[simp]"} attribute is exceptionally omitted for datatypes equipped +The \[simp]\ attribute is exceptionally omitted for datatypes equipped with a single nullary constructor, because a property of the form @{prop "x = C"} is not suitable as a simplification rule. -\item[@{text "t."}\hthm{distinct_disc} @{text "[dest]"}\rm:] ~ \\ +\item[\t.\\hthm{distinct_disc} \[dest]\\rm:] ~ \\ These properties are missing for @{typ "'a list"} because there is only one proper discriminator. If the datatype had been introduced with a second discriminator called @{const nonnull}, they would have read as follows: \\[\jot] @{prop "null list \ \ nonnull list"} \\ @{prop "nonnull list \ \ null list"} -\item[@{text "t."}\hthm{exhaust_disc} @{text "[case_names C\<^sub>1 \ C\<^sub>n]"}\rm:] ~ \\ +\item[\t.\\hthm{exhaust_disc} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust_disc[no_vars]} -\item[@{text "t."}\hthm{exhaust_sel} @{text "[case_names C\<^sub>1 \ C\<^sub>n]"}\rm:] ~ \\ +\item[\t.\\hthm{exhaust_sel} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust_sel[no_vars]} -\item[@{text "t."}\hthm{expand}\rm:] ~ \\ +\item[\t.\\hthm{expand}\rm:] ~ \\ @{thm list.expand[no_vars]} -\item[@{text "t."}\hthm{split_sel}\rm:] ~ \\ +\item[\t.\\hthm{split_sel}\rm:] ~ \\ @{thm list.split_sel[no_vars]} -\item[@{text "t."}\hthm{split_sel_asm}\rm:] ~ \\ +\item[\t.\\hthm{split_sel_asm}\rm:] ~ \\ @{thm list.split_sel_asm[no_vars]} -\item[@{text "t."}\hthm{split_sels} = @{text "split_sel split_sel_asm"}] - -\item[@{text "t."}\hthm{case_eq_if}\rm:] ~ \\ +\item[\t.\\hthm{split_sels} = \split_sel split_sel_asm\] + +\item[\t.\\hthm{case_eq_if}\rm:] ~ \\ @{thm list.case_eq_if[no_vars]} -\item[@{text "t."}\hthm{disc_eq_case}\rm:] ~ \\ +\item[\t.\\hthm{disc_eq_case}\rm:] ~ \\ @{thm list.disc_eq_case(1)[no_vars]} \\ @{thm list.disc_eq_case(2)[no_vars]} @@ -840,9 +840,9 @@ \end{indentblock} \noindent -In addition, equational versions of @{text t.disc} are registered with the -@{text "[code]"} attribute. The @{text "[code]"} attribute is set by the -@{text code} plugin (Section~\ref{ssec:code-generator}). +In addition, equational versions of \t.disc\ are registered with the +\[code]\ attribute. The \[code]\ attribute is set by the +\code\ plugin (Section~\ref{ssec:code-generator}). \ @@ -859,88 +859,87 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{case_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{case_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.case_transfer[no_vars]} \\ -This property is generated by the @{text transfer} plugin +This property is generated by the \transfer\ plugin (Section~\ref{ssec:transfer}). -%The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +%The \[transfer_rule]\ attribute is set by the \transfer\ plugin %(Section~\ref{ssec:transfer}). -\item[@{text "t."}\hthm{sel_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{sel_transfer} \[transfer_rule]\\rm:] ~ \\ This property is missing for @{typ "'a list"} because there is no common selector to all constructors. \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}). -\item[@{text "t."}\hthm{ctr_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{ctr_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.ctr_transfer(1)[no_vars]} \\ @{thm list.ctr_transfer(2)[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) . -\item[@{text "t."}\hthm{disc_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{disc_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.disc_transfer(1)[no_vars]} \\ @{thm list.disc_transfer(2)[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}). -\item[@{text "t."}\hthm{set} @{text "[simp, code]"}\rm:] ~ \\ +\item[\t.\\hthm{set} \[simp, code]\\rm:] ~ \\ @{thm list.set(1)[no_vars]} \\ @{thm list.set(2)[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "t."}\hthm{set_cases} @{text "[consumes 1, cases set: set\<^sub>i_t]"}\rm:] ~ \\ +\item[\t.\\hthm{set_cases} \[consumes 1, cases set: set\<^sub>i_t]\\rm:] ~ \\ @{thm list.set_cases[no_vars]} -\item[@{text "t."}\hthm{set_intros}\rm:] ~ \\ +\item[\t.\\hthm{set_intros}\rm:] ~ \\ @{thm list.set_intros(1)[no_vars]} \\ @{thm list.set_intros(2)[no_vars]} -\item[@{text "t."}\hthm{set_sel}\rm:] ~ \\ +\item[\t.\\hthm{set_sel}\rm:] ~ \\ @{thm list.set_sel[no_vars]} -\item[@{text "t."}\hthm{map} @{text "[simp, code]"}\rm:] ~ \\ +\item[\t.\\hthm{map} \[simp, code]\\rm:] ~ \\ @{thm list.map(1)[no_vars]} \\ @{thm list.map(2)[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "t."}\hthm{map_disc_iff} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{map_disc_iff} \[simp]\\rm:] ~ \\ @{thm list.map_disc_iff[no_vars]} -\item[@{text "t."}\hthm{map_sel}\rm:] ~ \\ +\item[\t.\\hthm{map_sel}\rm:] ~ \\ @{thm list.map_sel[no_vars]} -\item[@{text "t."}\hthm{pred_inject} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{pred_inject} \[simp]\\rm:] ~ \\ @{thm list.pred_inject(1)[no_vars]} \\ @{thm list.pred_inject(2)[no_vars]} -\item[@{text "t."}\hthm{rel_inject} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_inject} \[simp]\\rm:] ~ \\ @{thm list.rel_inject(1)[no_vars]} \\ @{thm list.rel_inject(2)[no_vars]} -\item[@{text "t."}\hthm{rel_distinct} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_distinct} \[simp]\\rm:] ~ \\ @{thm list.rel_distinct(1)[no_vars]} \\ @{thm list.rel_distinct(2)[no_vars]} -\item[@{text "t."}\hthm{rel_intros}\rm:] ~ \\ +\item[\t.\\hthm{rel_intros}\rm:] ~ \\ @{thm list.rel_intros(1)[no_vars]} \\ @{thm list.rel_intros(2)[no_vars]} -\item[@{text "t."}\hthm{rel_cases} @{text "[consumes 1, case_names t\<^sub>1 \ t\<^sub>m, cases pred]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_cases} \[consumes 1, case_names t\<^sub>1 \ t\<^sub>m, cases pred]\\rm:] ~ \\ @{thm list.rel_cases[no_vars]} -\item[@{text "t."}\hthm{rel_sel}\rm:] ~ \\ +\item[\t.\\hthm{rel_sel}\rm:] ~ \\ @{thm list.rel_sel[no_vars]} \end{description} \end{indentblock} \noindent -In addition, equational versions of @{text t.rel_inject} and @{text -rel_distinct} are registered with the @{text "[code]"} attribute. The -@{text "[code]"} attribute is set by the @{text code} plugin +In addition, equational versions of \t.rel_inject\ and \rel_distinct\ are registered with the \[code]\ attribute. The +\[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). The second subgroup consists of more abstract properties of the set functions, @@ -949,128 +948,128 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{inj_map}\rm:] ~ \\ +\item[\t.\\hthm{inj_map}\rm:] ~ \\ @{thm list.inj_map[no_vars]} -\item[@{text "t."}\hthm{inj_map_strong}\rm:] ~ \\ +\item[\t.\\hthm{inj_map_strong}\rm:] ~ \\ @{thm list.inj_map_strong[no_vars]} -\item[@{text "t."}\hthm{map_comp}\rm:] ~ \\ +\item[\t.\\hthm{map_comp}\rm:] ~ \\ @{thm list.map_comp[no_vars]} -\item[@{text "t."}\hthm{map_cong0}\rm:] ~ \\ +\item[\t.\\hthm{map_cong0}\rm:] ~ \\ @{thm list.map_cong0[no_vars]} -\item[@{text "t."}\hthm{map_cong} @{text "[fundef_cong]"}\rm:] ~ \\ +\item[\t.\\hthm{map_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.map_cong[no_vars]} -\item[@{text "t."}\hthm{map_cong_pred}\rm:] ~ \\ +\item[\t.\\hthm{map_cong_pred}\rm:] ~ \\ @{thm list.map_cong_pred[no_vars]} -\item[@{text "t."}\hthm{map_cong_simp}\rm:] ~ \\ +\item[\t.\\hthm{map_cong_simp}\rm:] ~ \\ @{thm list.map_cong_simp[no_vars]} -\item[@{text "t."}\hthm{map_id0}\rm:] ~ \\ +\item[\t.\\hthm{map_id0}\rm:] ~ \\ @{thm list.map_id0[no_vars]} -\item[@{text "t."}\hthm{map_id}\rm:] ~ \\ +\item[\t.\\hthm{map_id}\rm:] ~ \\ @{thm list.map_id[no_vars]} -\item[@{text "t."}\hthm{map_ident}\rm:] ~ \\ +\item[\t.\\hthm{map_ident}\rm:] ~ \\ @{thm list.map_ident[no_vars]} -\item[@{text "t."}\hthm{map_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{map_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.map_transfer[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. -\item[@{text "t."}\hthm{pred_cong} @{text "[fundef_cong]"}\rm:] ~ \\ +\item[\t.\\hthm{pred_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.pred_cong[no_vars]} -\item[@{text "t."}\hthm{pred_cong_simp}\rm:] ~ \\ +\item[\t.\\hthm{pred_cong_simp}\rm:] ~ \\ @{thm list.pred_cong_simp[no_vars]} -\item[@{text "t."}\hthm{pred_map}\rm:] ~ \\ +\item[\t.\\hthm{pred_map}\rm:] ~ \\ @{thm list.pred_map[no_vars]} -\item[@{text "t."}\hthm{pred_mono_strong}\rm:] ~ \\ +\item[\t.\\hthm{pred_mono_strong}\rm:] ~ \\ @{thm list.pred_mono_strong[no_vars]} -\item[@{text "t."}\hthm{pred_rel}\rm:] ~ \\ +\item[\t.\\hthm{pred_rel}\rm:] ~ \\ @{thm list.pred_rel[no_vars]} -\item[@{text "t."}\hthm{pred_set}\rm:] ~ \\ +\item[\t.\\hthm{pred_set}\rm:] ~ \\ @{thm list.pred_set[no_vars]} -\item[@{text "t."}\hthm{pred_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{pred_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.pred_transfer[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. -\item[@{text "t."}\hthm{pred_True}\rm:] ~ \\ +\item[\t.\\hthm{pred_True}\rm:] ~ \\ @{thm list.pred_True[no_vars]} -\item[@{text "t."}\hthm{set_map}\rm:] ~ \\ +\item[\t.\\hthm{set_map}\rm:] ~ \\ @{thm list.set_map[no_vars]} -\item[@{text "t."}\hthm{set_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{set_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.set_transfer[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. -\item[@{text "t."}\hthm{rel_compp} @{text "[relator_distr]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_compp} \[relator_distr]\\rm:] ~ \\ @{thm list.rel_compp[no_vars]} \\ -The @{text "[relator_distr]"} attribute is set by the @{text lifting} plugin +The \[relator_distr]\ attribute is set by the \lifting\ plugin (Section~\ref{ssec:lifting}). -\item[@{text "t."}\hthm{rel_conversep}\rm:] ~ \\ +\item[\t.\\hthm{rel_conversep}\rm:] ~ \\ @{thm list.rel_conversep[no_vars]} -\item[@{text "t."}\hthm{rel_eq}\rm:] ~ \\ +\item[\t.\\hthm{rel_eq}\rm:] ~ \\ @{thm list.rel_eq[no_vars]} -\item[@{text "t."}\hthm{rel_eq_onp}\rm:] ~ \\ +\item[\t.\\hthm{rel_eq_onp}\rm:] ~ \\ @{thm list.rel_eq_onp[no_vars]} -\item[@{text "t."}\hthm{rel_flip}\rm:] ~ \\ +\item[\t.\\hthm{rel_flip}\rm:] ~ \\ @{thm list.rel_flip[no_vars]} -\item[@{text "t."}\hthm{rel_map}\rm:] ~ \\ +\item[\t.\\hthm{rel_map}\rm:] ~ \\ @{thm list.rel_map(1)[no_vars]} \\ @{thm list.rel_map(2)[no_vars]} -\item[@{text "t."}\hthm{rel_mono} @{text "[mono, relator_mono]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_mono} \[mono, relator_mono]\\rm:] ~ \\ @{thm list.rel_mono[no_vars]} \\ -The @{text "[relator_mono]"} attribute is set by the @{text lifting} plugin +The \[relator_mono]\ attribute is set by the \lifting\ plugin (Section~\ref{ssec:lifting}). -\item[@{text "t."}\hthm{rel_mono_strong}\rm:] ~ \\ +\item[\t.\\hthm{rel_mono_strong}\rm:] ~ \\ @{thm list.rel_mono_strong[no_vars]} -\item[@{text "t."}\hthm{rel_cong} @{text "[fundef_cong]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.rel_cong[no_vars]} -\item[@{text "t."}\hthm{rel_cong_simp}\rm:] ~ \\ +\item[\t.\\hthm{rel_cong_simp}\rm:] ~ \\ @{thm list.rel_cong_simp[no_vars]} -\item[@{text "t."}\hthm{rel_refl}\rm:] ~ \\ +\item[\t.\\hthm{rel_refl}\rm:] ~ \\ @{thm list.rel_refl[no_vars]} -\item[@{text "t."}\hthm{rel_refl_strong}\rm:] ~ \\ +\item[\t.\\hthm{rel_refl_strong}\rm:] ~ \\ @{thm list.rel_refl_strong[no_vars]} -\item[@{text "t."}\hthm{rel_reflp}\rm:] ~ \\ +\item[\t.\\hthm{rel_reflp}\rm:] ~ \\ @{thm list.rel_reflp[no_vars]} -\item[@{text "t."}\hthm{rel_symp}\rm:] ~ \\ +\item[\t.\\hthm{rel_symp}\rm:] ~ \\ @{thm list.rel_symp[no_vars]} -\item[@{text "t."}\hthm{rel_transp}\rm:] ~ \\ +\item[\t.\\hthm{rel_transp}\rm:] ~ \\ @{thm list.rel_transp[no_vars]} -\item[@{text "t."}\hthm{rel_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.rel_transfer[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} @@ -1087,31 +1086,31 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n, induct t]"}\rm:] ~ \\ +\item[\t.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n, induct t]\\rm:] ~ \\ @{thm list.induct[no_vars]} -\item[@{text "t."}\hthm{rel_induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n, induct pred]"}\rm:] ~ \\ +\item[\t.\\hthm{rel_induct} \[case_names C\<^sub>1 \ C\<^sub>n, induct pred]\\rm:] ~ \\ @{thm list.rel_induct[no_vars]} \item[\begin{tabular}{@ {}l@ {}} - @{text "t\<^sub>1_\_t\<^sub>m."}\hthm{induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n]"}\rm: \\ - @{text "t\<^sub>1_\_t\<^sub>m."}\hthm{rel_induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n]"}\rm: \\ + \t\<^sub>1_\_t\<^sub>m.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm: \\ + \t\<^sub>1_\_t\<^sub>m.\\hthm{rel_induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm: \\ \end{tabular}] ~ \\ Given $m > 1$ mutually recursive datatypes, this induction rule can be used to prove $m$ properties simultaneously. -\item[@{text "t."}\hthm{rec} @{text "[simp, code]"}\rm:] ~ \\ +\item[\t.\\hthm{rec} \[simp, code]\\rm:] ~ \\ @{thm list.rec(1)[no_vars]} \\ @{thm list.rec(2)[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "t."}\hthm{rec_o_map}\rm:] ~ \\ +\item[\t.\\hthm{rec_o_map}\rm:] ~ \\ @{thm list.rec_o_map[no_vars]} -\item[@{text "t."}\hthm{rec_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{rec_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.rec_transfer[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} @@ -1123,8 +1122,8 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{simps}] = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.rec} @{text t.map} @{text t.rel_inject} \\ -@{text t.rel_distinct} @{text t.set} +\item[\t.\\hthm{simps}] = \t.inject\ \t.distinct\ \t.case\ \t.rec\ \t.map\ \t.rel_inject\ \\ +\t.rel_distinct\ \t.set\ \end{description} \end{indentblock} @@ -1161,10 +1160,10 @@ \item \emph{The Standard ML interfaces are different.} Tools and extensions written to call the old ML interfaces will need to be adapted to the new -interfaces. The @{text BNF_LFP_Compat} structure provides convenience functions +interfaces. The \BNF_LFP_Compat\ structure provides convenience functions that simulate the old interfaces in terms of the new ones. -\item \emph{The recursor @{text rec_t} has a different signature for nested +\item \emph{The recursor \rec_t\ has a different signature for nested recursive datatypes.} In the old package, nested recursion through non-functions was internally reduced to mutual recursion. This reduction was visible in the type of the recursor, used by \keyw{primrec}. Recursion through functions was @@ -1180,17 +1179,17 @@ using @{command primrec} if the recursion is via new-style datatypes, as explained in Section~\ref{sssec:primrec-nested-as-mutual-recursion}, or using @{command datatype_compat}. For recursion through functions, the old-style -induction rule can be obtained by applying the @{text "[unfolded -all_mem_range]"} attribute on @{text t.induct}. +induction rule can be obtained by applying the \[unfolded +all_mem_range]\ attribute on \t.induct\. \item \emph{The @{const size} function has a slightly different definition.} -The new function returns @{text 1} instead of @{text 0} for some nonrecursive +The new function returns \1\ instead of \0\ for some nonrecursive constructors. This departure from the old behavior made it possible to implement -@{const size} in terms of the generic function @{text "t.size_t"}. Moreover, +@{const size} in terms of the generic function \t.size_t\. Moreover, the new function considers nested occurrences of a value, in the nested -recursive case. The old behavior can be obtained by disabling the @{text size} +recursive case. The old behavior can be obtained by disabling the \size\ plugin (Section~\ref{sec:selecting-plugins}) and instantiating the -@{text size} type class manually. +\size\ type class manually. \item \emph{The internal constructions are completely different.} Proof texts that unfold the definition of constants introduced by the old command will @@ -1198,23 +1197,23 @@ \item \emph{Some constants and theorems have different names.} For non-mutually recursive datatypes, -the alias @{text t.inducts} for @{text t.induct} is no longer generated. +the alias \t.inducts\ for \t.induct\ is no longer generated. For $m > 1$ mutually recursive datatypes, -@{text "rec_t\<^sub>1_\_t\<^sub>m_i"} has been renamed -@{text "rec_t\<^sub>i"} for each @{text "i \ {1, \, m}"}, -@{text "t\<^sub>1_\_t\<^sub>m.inducts(i)"} has been renamed -@{text "t\<^sub>i.induct"} for each @{text "i \ {1, \, m}"}, and the -collection @{text "t\<^sub>1_\_t\<^sub>m.size"} (generated by the -@{text size} plugin, Section~\ref{ssec:size}) has been divided into -@{text "t\<^sub>1.size"}, \ldots, @{text "t\<^sub>m.size"}. - -\item \emph{The @{text t.simps} collection has been extended.} +\rec_t\<^sub>1_\_t\<^sub>m_i\ has been renamed +\rec_t\<^sub>i\ for each \i \ {1, \, m}\, +\t\<^sub>1_\_t\<^sub>m.inducts(i)\ has been renamed +\t\<^sub>i.induct\ for each \i \ {1, \, m}\, and the +collection \t\<^sub>1_\_t\<^sub>m.size\ (generated by the +\size\ plugin, Section~\ref{ssec:size}) has been divided into +\t\<^sub>1.size\, \ldots, \t\<^sub>m.size\. + +\item \emph{The \t.simps\ collection has been extended.} Previously available theorems are available at the same index as before. \item \emph{Variables in generated properties have different names.} This is rarely an issue, except in proof texts that refer to variable names in the -@{text "[where \]"} attribute. The solution is to use the more robust -@{text "[of \]"} syntax. +\[where \]\ attribute. The solution is to use the more robust +\[of \]\ syntax. \end{itemize} \ @@ -1396,8 +1395,8 @@ text \ \noindent -The next example features recursion through the @{text option} type. Although -@{text option} is not a new-style datatype, it is registered as a BNF with the +The next example features recursion through the \option\ type. Although +\option\ is not a new-style datatype, it is registered as a BNF with the map function @{const map_option}: \ @@ -1410,7 +1409,7 @@ \noindent The same principle applies for arbitrary type constructors through which recursion is possible. Notably, the map function for the function type -(@{text \}) is simply composition (@{text "(\)"}): +(\\\) is simply composition (\(\)\): \ primrec (*<*)(in early) (*>*)relabel_ft :: "('a \ 'a) \ 'a ftree \ 'a ftree" where @@ -1559,7 +1558,7 @@ text \ \begin{matharray}{rcl} - @{command_def "primrec"} & : & @{text "local_theory \ local_theory"} + @{command_def "primrec"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -1589,18 +1588,18 @@ \setlength{\itemsep}{0pt} \item -The @{text plugins} option indicates which plugins should be enabled -(@{text only}) or disabled (@{text del}). By default, all plugins are enabled. +The \plugins\ option indicates which plugins should be enabled +(\only\) or disabled (\del\). By default, all plugins are enabled. \item -The @{text nonexhaustive} option indicates that the functions are not +The \nonexhaustive\ option indicates that the functions are not necessarily specified for all constructors. It can be used to suppress the warning that is normally emitted when some constructors are missing. \item -The @{text transfer} option indicates that an unconditional transfer rule -should be generated and proved @{text "by transfer_prover"}. The -@{text "[transfer_rule]"} attribute is set on the generated theorem. +The \transfer\ option indicates that an unconditional transfer rule +should be generated and proved \by transfer_prover\. The +\[transfer_rule]\ attribute is set on the generated theorem. \end{itemize} %%% TODO: elaborate on format of the equations @@ -1623,23 +1622,23 @@ \begin{indentblock} \begin{description} -\item[@{text "f."}\hthm{simps} @{text "[simp, code]"}\rm:] ~ \\ +\item[\f.\\hthm{simps} \[simp, code]\\rm:] ~ \\ @{thm tfold.simps(1)[no_vars]} \\ @{thm tfold.simps(2)[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "f."}\hthm{transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\f.\\hthm{transfer} \[transfer_rule]\\rm:] ~ \\ @{thm tfold.transfer[no_vars]} \\ -This theorem is generated by the @{text transfer} plugin -(Section~\ref{ssec:transfer}) for functions declared with the @{text transfer} +This theorem is generated by the \transfer\ plugin +(Section~\ref{ssec:transfer}) for functions declared with the \transfer\ option enabled. -\item[@{text "f."}\hthm{induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n]"}\rm:] ~ \\ +\item[\f.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ This induction rule is generated for nested-as-mutual recursive functions (Section~\ref{sssec:primrec-nested-as-mutual-recursion}). -\item[@{text "f\<^sub>1_\_f\<^sub>m."}\hthm{induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n]"}\rm:] ~ \\ +\item[\f\<^sub>1_\_f\<^sub>m.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ This induction rule is generated for nested-as-mutual recursive functions (Section~\ref{sssec:primrec-nested-as-mutual-recursion}). Given $m > 1$ mutually recursive functions, this rule can be used to prove $m$ properties @@ -1658,7 +1657,7 @@ \label{ssec:primrec-recursive-default-values-for-selectors}\ text \ -A datatype selector @{text un_D} can have a default value for each constructor +A datatype selector \un_D\ can have a default value for each constructor on which it is not otherwise specified. Occasionally, it is useful to have the default value be defined recursively. This leads to a chicken-and-egg situation, because the datatype is not introduced yet at the moment when the @@ -1673,20 +1672,20 @@ \setlength{\itemsep}{0pt} \item -Introduce a fully unspecified constant @{text "un_D\<^sub>0 :: 'a"} using +Introduce a fully unspecified constant \un_D\<^sub>0 :: 'a\ using @{command consts}. \item -Define the datatype, specifying @{text "un_D\<^sub>0"} as the selector's default +Define the datatype, specifying \un_D\<^sub>0\ as the selector's default value. \item -Define the behavior of @{text "un_D\<^sub>0"} on values of the newly introduced +Define the behavior of \un_D\<^sub>0\ on values of the newly introduced datatype using the \keyw{overloading} command. \item -Derive the desired equation on @{text un_D} from the characteristic equations -for @{text "un_D\<^sub>0"}. +Derive the desired equation on \un_D\ from the characteristic equations +for \un_D\<^sub>0\. \end{enumerate} \noindent @@ -1737,8 +1736,8 @@ \item \emph{Some theorems have different names.} For $m > 1$ mutually recursive functions, -@{text "f\<^sub>1_\_f\<^sub>m.simps"} has been broken down into separate -subcollections @{text "f\<^sub>i.simps"}. +\f\<^sub>1_\_f\<^sub>m.simps\ has been broken down into separate +subcollections \f\<^sub>i.simps\. \end{itemize} \ @@ -1781,8 +1780,8 @@ text \ \noindent -Lazy lists can be infinite, such as @{text "LCons 0 (LCons 0 (\))"} and -@{text "LCons 0 (LCons 1 (LCons 2 (\)))"}. Here is a related type, that of +Lazy lists can be infinite, such as \LCons 0 (LCons 0 (\))\ and +\LCons 0 (LCons 1 (LCons 2 (\)))\. Here is a related type, that of infinite streams: \ @@ -1802,9 +1801,9 @@ text \ \noindent -This type has exactly one infinite element, @{text "ESucc (ESucc (ESucc (\)))"}, +This type has exactly one infinite element, \ESucc (ESucc (ESucc (\)))\, that represents $\infty$. In addition, it has finite values of the form -@{text "ESucc (\ (ESucc EZero)\)"}. +\ESucc (\ (ESucc EZero)\)\. Here is an example with many constructors: \ @@ -1861,7 +1860,7 @@ text \ \begin{matharray}{rcl} - @{command_def "codatatype"} & : & @{text "local_theory \ local_theory"} + @{command_def "codatatype"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -1872,7 +1871,7 @@ \noindent Definitions of codatatypes have almost exactly the same syntax as for datatypes -(Section~\ref{ssec:datatype-command-syntax}). The @{text "discs_sels"} option +(Section~\ref{ssec:datatype-command-syntax}). The \discs_sels\ option is superfluous because discriminators and selectors are always generated for codatatypes. \ @@ -1882,9 +1881,9 @@ \label{ssec:codatatype-generated-constants}\ text \ -Given a codatatype @{text "('a\<^sub>1, \, 'a\<^sub>m) t"} -with $m > 0$ live type variables and $n$ constructors @{text "t.C\<^sub>1"}, -\ldots, @{text "t.C\<^sub>n"}, the same auxiliary constants are generated as for +Given a codatatype \('a\<^sub>1, \, 'a\<^sub>m) t\ +with $m > 0$ live type variables and $n$ constructors \t.C\<^sub>1\, +\ldots, \t.C\<^sub>n\, the same auxiliary constants are generated as for datatypes (Section~\ref{ssec:datatype-generated-constants}), except that the recursor is replaced by a dual concept: @@ -1892,7 +1891,7 @@ \begin{tabular}{@ {}ll@ {}} Corecursor: & - @{text t.corec_t} + \t.corec_t\ \end{tabular} \ @@ -1934,69 +1933,69 @@ \begin{description} \item[\begin{tabular}{@ {}l@ {}} - @{text "t."}\hthm{coinduct} @{text "[consumes m, case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "t."}\hthm{coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \ - D\<^sub>n, coinduct t]"}\rm: + \t.\\hthm{coinduct} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\t.\\hthm{coinduct} \[\}\case_conclusion D\<^sub>1 \ + D\<^sub>n, coinduct t]\\rm: \end{tabular}] ~ \\ @{thm llist.coinduct[no_vars]} \item[\begin{tabular}{@ {}l@ {}} - @{text "t."}\hthm{coinduct_strong} @{text "[consumes m, case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "t."}\hthm{coinduct_strong} @{text "["}}@{text "case_conclusion D\<^sub>1 \ D\<^sub>n]"}\rm: + \t.\\hthm{coinduct_strong} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\t.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ @{thm llist.coinduct_strong[no_vars]} \item[\begin{tabular}{@ {}l@ {}} - @{text "t."}\hthm{rel_coinduct} @{text "[consumes m, case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "t."}\hthm{rel_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \ - D\<^sub>n, coinduct pred]"}\rm: + \t.\\hthm{rel_coinduct} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\t.\\hthm{rel_coinduct} \[\}\case_conclusion D\<^sub>1 \ + D\<^sub>n, coinduct pred]\\rm: \end{tabular}] ~ \\ @{thm llist.rel_coinduct[no_vars]} \item[\begin{tabular}{@ {}l@ {}} - @{text "t\<^sub>1_\_t\<^sub>m."}\hthm{coinduct} @{text "[case_names t\<^sub>1 \ t\<^sub>m, case_conclusion D\<^sub>1 \ D\<^sub>n]"} \\ - @{text "t\<^sub>1_\_t\<^sub>m."}\hthm{coinduct_strong} @{text "[case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "t\<^sub>1_\_t\<^sub>m."}\hthm{coinduct_strong} @{text "["}}@{text "case_conclusion D\<^sub>1 \ D\<^sub>n]"}\rm: \\ - @{text "t\<^sub>1_\_t\<^sub>m."}\hthm{rel_coinduct} @{text "[case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "t\<^sub>1_\_t\<^sub>m."}\hthm{rel_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \ D\<^sub>n]"}\rm: \\ + \t\<^sub>1_\_t\<^sub>m.\\hthm{coinduct} \[case_names t\<^sub>1 \ t\<^sub>m, case_conclusion D\<^sub>1 \ D\<^sub>n]\ \\ + \t\<^sub>1_\_t\<^sub>m.\\hthm{coinduct_strong} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\t\<^sub>1_\_t\<^sub>m.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \\ + \t\<^sub>1_\_t\<^sub>m.\\hthm{rel_coinduct} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\t\<^sub>1_\_t\<^sub>m.\\hthm{rel_coinduct} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \\ \end{tabular}] ~ \\ Given $m > 1$ mutually corecursive codatatypes, these coinduction rules can be used to prove $m$ properties simultaneously. \item[\begin{tabular}{@ {}l@ {}} - @{text "t\<^sub>1_\_t\<^sub>m."}\hthm{set_induct} @{text "[case_names C\<^sub>1 \ C\<^sub>n,"} \\ - \phantom{@{text "t\<^sub>1_\_t\<^sub>m."}\hthm{set_induct} @{text "["}}@{text "induct set: set\<^sub>j_t\<^sub>1, \, induct set: set\<^sub>j_t\<^sub>m]"}\rm: \\ + \t\<^sub>1_\_t\<^sub>m.\\hthm{set_induct} \[case_names C\<^sub>1 \ C\<^sub>n,\ \\ + \phantom{\t\<^sub>1_\_t\<^sub>m.\\hthm{set_induct} \[\}\induct set: set\<^sub>j_t\<^sub>1, \, induct set: set\<^sub>j_t\<^sub>m]\\rm: \\ \end{tabular}] ~ \\ @{thm llist.set_induct[no_vars]} \\ -If $m = 1$, the attribute @{text "[consumes 1]"} is generated as well. - -\item[@{text "t."}\hthm{corec}\rm:] ~ \\ +If $m = 1$, the attribute \[consumes 1]\ is generated as well. + +\item[\t.\\hthm{corec}\rm:] ~ \\ @{thm llist.corec(1)[no_vars]} \\ @{thm llist.corec(2)[no_vars]} -\item[@{text "t."}\hthm{corec_code} @{text "[code]"}\rm:] ~ \\ +\item[\t.\\hthm{corec_code} \[code]\\rm:] ~ \\ @{thm llist.corec_code[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "t."}\hthm{corec_disc}\rm:] ~ \\ +\item[\t.\\hthm{corec_disc}\rm:] ~ \\ @{thm llist.corec_disc(1)[no_vars]} \\ @{thm llist.corec_disc(2)[no_vars]} -\item[@{text "t."}\hthm{corec_disc_iff} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{corec_disc_iff} \[simp]\\rm:] ~ \\ @{thm llist.corec_disc_iff(1)[no_vars]} \\ @{thm llist.corec_disc_iff(2)[no_vars]} -\item[@{text "t."}\hthm{corec_sel} @{text "[simp]"}\rm:] ~ \\ +\item[\t.\\hthm{corec_sel} \[simp]\\rm:] ~ \\ @{thm llist.corec_sel(1)[no_vars]} \\ @{thm llist.corec_sel(2)[no_vars]} -\item[@{text "t."}\hthm{map_o_corec}\rm:] ~ \\ +\item[\t.\\hthm{map_o_corec}\rm:] ~ \\ @{thm llist.map_o_corec[no_vars]} -\item[@{text "t."}\hthm{corec_transfer} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{corec_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm llist.corec_transfer[no_vars]} \\ -The @{text "[transfer_rule]"} attribute is set by the @{text transfer} plugin +The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} @@ -2008,8 +2007,8 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{simps}] = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.corec_disc_iff} @{text t.corec_sel} \\ -@{text t.map} @{text t.rel_inject} @{text t.rel_distinct} @{text t.set} +\item[\t.\\hthm{simps}] = \t.inject\ \t.distinct\ \t.case\ \t.corec_disc_iff\ \t.corec_sel\ \\ +\t.map\ \t.rel_inject\ \t.rel_distinct\ \t.set\ \end{description} \end{indentblock} @@ -2042,17 +2041,17 @@ \belowdisplayskip=.5\belowdisplayskip \item The \emph{destructor view} specifies $f$ by implications of the form -\[@{text "\ \ is_C\<^sub>j (f x\<^sub>1 \ x\<^sub>n)"}\] and +\[\\ \ is_C\<^sub>j (f x\<^sub>1 \ x\<^sub>n)\\] and equations of the form -\[@{text "un_C\<^sub>ji (f x\<^sub>1 \ x\<^sub>n) = \"}\] +\[\un_C\<^sub>ji (f x\<^sub>1 \ x\<^sub>n) = \\\] This style is popular in the coalgebraic literature. \item The \emph{constructor view} specifies $f$ by equations of the form -\[@{text "\ \ f x\<^sub>1 \ x\<^sub>n = C\<^sub>j \"}\] +\[\\ \ f x\<^sub>1 \ x\<^sub>n = C\<^sub>j \\\] This style is often more concise than the previous one. \item The \emph{code view} specifies $f$ by a single equation of the form -\[@{text "f x\<^sub>1 \ x\<^sub>n = \"}\] +\[\f x\<^sub>1 \ x\<^sub>n = \\\] with restrictions on the format of the right-hand side. Lazy functional programming languages such as Haskell support a generalized version of this style. @@ -2100,9 +2099,9 @@ \noindent The constructor ensures that progress is made---i.e., the function is \emph{productive}. The above functions compute the infinite lazy list or stream -@{text "[x, g x, g (g x), \]"}. Productivity guarantees that prefixes -@{text "[x, g x, g (g x), \, (g ^^ k) x]"} of arbitrary finite length -@{text k} can be computed by unfolding the code equation a finite number of +\[x, g x, g (g x), \]\. Productivity guarantees that prefixes +\[x, g x, g (g x), \, (g ^^ k) x]\ of arbitrary finite length +\k\ can be computed by unfolding the code equation a finite number of times. Corecursive functions construct codatatype values, but nothing prevents them @@ -2115,8 +2114,7 @@ text \ \noindent -Constructs such as @{text "let"}--@{text "in"}, @{text -"if"}--@{text "then"}--@{text "else"}, and @{text "case"}--@{text "of"} may +Constructs such as \let\--\in\, \if\--\then\--\else\, and \case\--\of\ may appear around constructors that guard corecursive calls: \ @@ -2128,7 +2126,7 @@ text \ \noindent -For technical reasons, @{text "case"}--@{text "of"} is only supported for +For technical reasons, \case\--\of\ is only supported for case distinctions on (co)datatypes that provide discriminators and selectors. Pattern matching is not supported by @{command primcorec}. Fortunately, it is @@ -2155,8 +2153,8 @@ text \ \noindent The example below constructs a pseudorandom process value. It takes a stream of -actions (@{text s}), a pseudorandom function generator (@{text f}), and a -pseudorandom seed (@{text n}): +actions (\s\), a pseudorandom function generator (\f\), and a +pseudorandom seed (\n\): \ (*<*) @@ -2210,8 +2208,7 @@ text \ The next pair of examples generalize the @{const literate} and @{const siterate} functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly -infinite trees in which subnodes are organized either as a lazy list (@{text -tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}). They rely on the map functions of +infinite trees in which subnodes are organized either as a lazy list (\tree\<^sub>i\<^sub>i\) or as a finite set (\tree\<^sub>i\<^sub>s\). They rely on the map functions of the nesting type constructors to lift the corecursive calls: \ @@ -2252,9 +2249,9 @@ text \ The next example illustrates corecursion through functions, which is a bit special. Deterministic finite automata (DFAs) are traditionally defined as -5-tuples @{text "(Q, \, \, q\<^sub>0, F)"}, where @{text Q} is a finite set of states, -@{text \} is a finite alphabet, @{text \} is a transition function, @{text q\<^sub>0} -is an initial state, and @{text F} is a set of final states. The following +5-tuples \(Q, \, \, q\<^sub>0, F)\, where \Q\ is a finite set of states, +\\\ is a finite alphabet, \\\ is a transition function, \q\<^sub>0\ +is an initial state, and \F\ is a set of final states. The following function translates a DFA into a state machine: \ @@ -2263,8 +2260,8 @@ text \ \noindent -The map function for the function type (@{text \}) is composition -(@{text "(\)"}). For convenience, corecursion through functions can +The map function for the function type (\\\) is composition +(\(\)\). For convenience, corecursion through functions can also be expressed using $\lambda$-abstractions and function application rather than through composition. For example: \ @@ -2343,7 +2340,7 @@ @{thm [source] iterate\<^sub>i\<^sub>i.coinduct}, @{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and @{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct} -and analogously for @{text coinduct_strong}. These rules and the +and analogously for \coinduct_strong\. These rules and the underlying corecursors are generated dynamically and are kept in a cache to speed up subsequent definitions. \ @@ -2418,7 +2415,7 @@ obligations, one for each pair of conditions @{term "(n mod (4::int) = i, n mod (4::int) = j)"} with @{term "i < j"}. If we prefer not to discharge any obligations, we can -enable the @{text "sequential"} option. This pushes the problem to the users of +enable the \sequential\ option. This pushes the problem to the users of the generated properties. %Here are more examples to conclude: \ @@ -2435,7 +2432,7 @@ text \ The destructor view is in many respects dual to the constructor view. Conditions determine which constructor to choose, and these conditions are interpreted -sequentially or not depending on the @{text "sequential"} option. +sequentially or not depending on the \sequential\ option. Consider the following examples: \ @@ -2517,7 +2514,7 @@ text \ \noindent -Using the @{text "of"} keyword, different equations are specified for @{const +Using the \of\ keyword, different equations are specified for @{const cont} depending on which constructor is selected. Here are more examples to conclude: @@ -2549,8 +2546,8 @@ text \ \begin{matharray}{rcl} - @{command_def "primcorec"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "primcorecursive"} & : & @{text "local_theory \ proof(prove)"} + @{command_def "primcorec"} & : & \local_theory \ local_theory\ \\ + @{command_def "primcorecursive"} & : & \local_theory \ proof(prove)\ \end{matharray} @{rail \ @@ -2580,27 +2577,27 @@ \setlength{\itemsep}{0pt} \item -The @{text plugins} option indicates which plugins should be enabled -(@{text only}) or disabled (@{text del}). By default, all plugins are enabled. +The \plugins\ option indicates which plugins should be enabled +(\only\) or disabled (\del\). By default, all plugins are enabled. \item -The @{text sequential} option indicates that the conditions in specifications +The \sequential\ option indicates that the conditions in specifications expressed using the constructor or destructor view are to be interpreted sequentially. \item -The @{text exhaustive} option indicates that the conditions in specifications +The \exhaustive\ option indicates that the conditions in specifications expressed using the constructor or destructor view cover all possible cases. This generally gives rise to an additional proof obligation. \item -The @{text transfer} option indicates that an unconditional transfer rule -should be generated and proved @{text "by transfer_prover"}. The -@{text "[transfer_rule]"} attribute is set on the generated theorem. +The \transfer\ option indicates that an unconditional transfer rule +should be generated and proved \by transfer_prover\. The +\[transfer_rule]\ attribute is set on the generated theorem. \end{itemize} The @{command primcorec} command is an abbreviation for @{command -primcorecursive} with @{text "by auto?"} to discharge any emerging proof +primcorecursive} with \by auto?\ to discharge any emerging proof obligations. %%% TODO: elaborate on format of the propositions @@ -2618,60 +2615,60 @@ \begin{indentblock} \begin{description} -\item[@{text "f."}\hthm{code} @{text "[code]"}\rm:] ~ \\ +\item[\f.\\hthm{code} \[code]\\rm:] ~ \\ @{thm literate.code[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "f."}\hthm{ctr}\rm:] ~ \\ +\item[\f.\\hthm{ctr}\rm:] ~ \\ @{thm literate.ctr[no_vars]} -\item[@{text "f."}\hthm{disc} @{text "[simp, code]"}\rm:] ~ \\ +\item[\f.\\hthm{disc} \[simp, code]\\rm:] ~ \\ @{thm literate.disc[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin -(Section~\ref{ssec:code-generator}). The @{text "[simp]"} attribute is set only -for functions for which @{text f.disc_iff} is not available. - -\item[@{text "f."}\hthm{disc_iff} @{text "[simp]"}\rm:] ~ \\ +The \[code]\ attribute is set by the \code\ plugin +(Section~\ref{ssec:code-generator}). The \[simp]\ attribute is set only +for functions for which \f.disc_iff\ is not available. + +\item[\f.\\hthm{disc_iff} \[simp]\\rm:] ~ \\ @{thm literate.disc_iff[no_vars]} \\ This property is generated only for functions declared with the -@{text exhaustive} option or whose conditions are trivially exhaustive. - -\item[@{text "f."}\hthm{sel} @{text "[simp, code]"}\rm:] ~ \\ +\exhaustive\ option or whose conditions are trivially exhaustive. + +\item[\f.\\hthm{sel} \[simp, code]\\rm:] ~ \\ @{thm literate.disc[no_vars]} \\ -The @{text "[code]"} attribute is set by the @{text code} plugin +The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). -\item[@{text "f."}\hthm{exclude}\rm:] ~ \\ +\item[\f.\\hthm{exclude}\rm:] ~ \\ These properties are missing for @{const literate} because no exclusiveness proof obligations arose. In general, the properties correspond to the discharged proof obligations. -\item[@{text "f."}\hthm{exhaust}\rm:] ~ \\ +\item[\f.\\hthm{exhaust}\rm:] ~ \\ This property is missing for @{const literate} because no exhaustiveness proof obligation arose. In general, the property correspond to the discharged proof obligation. \item[\begin{tabular}{@ {}l@ {}} - @{text "f."}\hthm{coinduct} @{text "[consumes m, case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "f."}\hthm{coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \ - D\<^sub>n]"}\rm: + \f.\\hthm{coinduct} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\f.\\hthm{coinduct} \[\}\case_conclusion D\<^sub>1 \ + D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). \item[\begin{tabular}{@ {}l@ {}} - @{text "f."}\hthm{coinduct_strong} @{text "[consumes m, case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "f."}\hthm{coinduct_strong} @{text "["}}@{text "case_conclusion D\<^sub>1 \ - D\<^sub>n]"}\rm: + \f.\\hthm{coinduct_strong} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\f.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ + D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). \item[\begin{tabular}{@ {}l@ {}} - @{text "f\<^sub>1_\_f\<^sub>m."}\hthm{coinduct} @{text "[case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "f."}\hthm{coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \ - D\<^sub>n]"}\rm: + \f\<^sub>1_\_f\<^sub>m.\\hthm{coinduct} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\f.\\hthm{coinduct} \[\}\case_conclusion D\<^sub>1 \ + D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). Given $m > 1$ @@ -2679,9 +2676,9 @@ simultaneously. \item[\begin{tabular}{@ {}l@ {}} - @{text "f\<^sub>1_\_f\<^sub>m."}\hthm{coinduct_strong} @{text "[case_names t\<^sub>1 \ t\<^sub>m,"} \\ - \phantom{@{text "f."}\hthm{coinduct_strong} @{text "["}}@{text "case_conclusion D\<^sub>1 \ - D\<^sub>n]"}\rm: + \f\<^sub>1_\_f\<^sub>m.\\hthm{coinduct_strong} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ + \phantom{\f.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ + D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). Given $m > 1$ @@ -2698,7 +2695,7 @@ \begin{indentblock} \begin{description} -\item[@{text "f."}\hthm{simps}] = @{text f.disc_iff} (or @{text f.disc}) @{text t.sel} +\item[\f.\\hthm{simps}] = \f.disc_iff\ (or \f.disc\) \t.sel\ \end{description} \end{indentblock} @@ -2738,14 +2735,14 @@ (functorial action), $n$ set functions (natural transformations), and an infinite cardinal bound that satisfy certain properties. For example, @{typ "'a llist"} is a unary BNF. -Its predicator @{text "llist_all :: +Its predicator \llist_all :: ('a \ bool) \ - 'a llist \ bool"} + 'a llist \ bool\ extends unary predicates over elements to unary predicates over lazy lists. -Similarly, its relator @{text "llist_all2 :: +Similarly, its relator \llist_all2 :: ('a \ 'b \ bool) \ - 'a llist \ 'b llist \ bool"} + 'a llist \ 'b llist \ bool\ extends binary predicates over elements to binary predicates over parallel lazy lists. The cardinal bound limits the number of elements returned by the set function; it may not depend on the cardinality of @{typ 'a}. @@ -2947,11 +2944,11 @@ text \ The next example declares a BNF axiomatically. This can be convenient for reasoning abstractly about an arbitrary BNF. The @{command bnf_axiomatization} -command below introduces a type @{text "('a, 'b, 'c) F"}, three set constants, +command below introduces a type \('a, 'b, 'c) F\, three set constants, a map function, a predicator, a relator, and a nonemptiness witness that depends only on -@{typ 'a}. The type @{text "'a \ ('a, 'b, 'c) F"} of the witness can be read +@{typ 'a}. The type \'a \ ('a, 'b, 'c) F\ of the witness can be read as an implication: Given a witness for @{typ 'a}, we can construct a witness for -@{text "('a, 'b, 'c) F"}. The BNF properties are postulated as axioms. +\('a, 'b, 'c) F\. The BNF properties are postulated as axioms. \ bnf_axiomatization (setA: 'a, setB: 'b, setC: 'c) F @@ -2971,7 +2968,7 @@ text \ \begin{matharray}{rcl} - @{command_def "bnf"} & : & @{text "local_theory \ proof(prove)"} + @{command_def "bnf"} & : & \local_theory \ proof(prove)\ \end{matharray} @{rail \ @@ -2994,7 +2991,7 @@ @{cite "isabelle-isar-ref"}. The @{syntax plugins} option indicates which plugins should be enabled -(@{text only}) or disabled (@{text del}). By default, all plugins are enabled. +(\only\) or disabled (\del\). By default, all plugins are enabled. %%% TODO: elaborate on proof obligations \ @@ -3004,7 +3001,7 @@ text \ \begin{matharray}{rcl} - @{command_def "lift_bnf"} & : & @{text "local_theory \ proof(prove)"} + @{command_def "lift_bnf"} & : & \local_theory \ proof(prove)\ \end{matharray} @{rail \ @@ -3024,7 +3021,7 @@ type}) using the @{command typedef} command. To achieve this, it lifts the BNF structure on the raw type to the abstract type following a @{term type_definition} theorem. The theorem is usually inferred from the type, but can -also be explicitly supplied by means of the optional @{text via} clause. In +also be explicitly supplied by means of the optional \via\ clause. In addition, custom names for the set functions, the map function, the predicator, and the relator, as well as nonemptiness witnesses can be specified. @@ -3032,7 +3029,7 @@ incomplete. They must be given as terms (on the raw type) and proved to be witnesses. The command warns about witness types that are present in the raw type's BNF but not supplied by the user. The warning can be disabled by -specifying the @{text no_warn_wits} option. +specifying the \no_warn_wits\ option. \ subsubsection \\keyw{copy_bnf} @@ -3040,7 +3037,7 @@ text \ \begin{matharray}{rcl} - @{command_def "copy_bnf"} & : & @{text "local_theory \ local_theory"} + @{command_def "copy_bnf"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -3061,7 +3058,7 @@ text \ \begin{matharray}{rcl} - @{command_def "bnf_axiomatization"} & : & @{text "local_theory \ local_theory"} + @{command_def "bnf_axiomatization"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -3086,10 +3083,10 @@ @{cite "isabelle-isar-ref"}. The @{syntax plugins} option indicates which plugins should be enabled -(@{text only}) or disabled (@{text del}). By default, all plugins are enabled. +(\only\) or disabled (\del\). By default, all plugins are enabled. Type arguments are live by default; they can be marked as dead by entering -@{text dead} in front of the type variable (e.g., @{text "(dead 'a)"}) +\dead\ in front of the type variable (e.g., \(dead 'a)\) instead of an identifier for the corresponding set function. Witnesses can be specified by their types. Otherwise, the syntax of @{command bnf_axiomatization} is identical to the left-hand side of a @{command datatype} or @@ -3107,7 +3104,7 @@ text \ \begin{matharray}{rcl} - @{command_def "print_bnfs"} & : & @{text "local_theory \"} + @{command_def "print_bnfs"} & : & \local_theory \\ \end{matharray} @{rail \ @@ -3128,7 +3125,7 @@ % a type not introduced by ... % % * @{command free_constructors} -% * @{text plugins}, @{text discs_sels} +% * \plugins\, \discs_sels\ % * hack to have both co and nonco view via locale (cf. ext nats) \ @@ -3147,7 +3144,7 @@ text \ \begin{matharray}{rcl} - @{command_def "free_constructors"} & : & @{text "local_theory \ proof(prove)"} + @{command_def "free_constructors"} & : & \local_theory \ proof(prove)\ \end{matharray} @{rail \ @@ -3177,8 +3174,8 @@ constructor itself (as a term), and a list of optional names for the selectors. Section~\ref{ssec:datatype-generated-theorems} lists the generated theorems. -For bootstrapping reasons, the generally useful @{text "[fundef_cong]"} -attribute is not set on the generated @{text case_cong} theorem. It can be +For bootstrapping reasons, the generally useful \[fundef_cong]\ +attribute is not set on the generated \case_cong\ theorem. It can be added manually using \keyw{declare}. \ @@ -3188,7 +3185,7 @@ text \ \begin{matharray}{rcl} - @{command_def "simps_of_case"} & : & @{text "local_theory \ local_theory"} + @{command_def "simps_of_case"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -3227,7 +3224,7 @@ text \ \begin{matharray}{rcl} - @{command_def "case_of_simps"} & : & @{text "local_theory \ local_theory"} + @{command_def "case_of_simps"} & : & \local_theory \ local_theory\ \end{matharray} @{rail \ @@ -3308,10 +3305,10 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{eq.refl} @{text "[code nbe]"}\rm:] ~ \\ +\item[\t.\\hthm{eq.refl} \[code nbe]\\rm:] ~ \\ @{thm list.eq.refl[no_vars]} -\item[@{text "t."}\hthm{eq.simps} @{text "[code]"}\rm:] ~ \\ +\item[\t.\\hthm{eq.simps} \[code]\\rm:] ~ \\ @{thm list.eq.simps(1)[no_vars]} \\ @{thm list.eq.simps(2)[no_vars]} \\ @{thm list.eq.simps(3)[no_vars]} \\ @@ -3322,7 +3319,7 @@ \end{description} \end{indentblock} -In addition, the plugin sets the @{text "[code]"} attribute on a number of +In addition, the plugin sets the \[code]\ attribute on a number of properties of freely generated types and of (co)recursive functions, as documented in Sections \ref{ssec:datatype-generated-theorems}, \ref{ssec:primrec-generated-theorems}, \ref{ssec:codatatype-generated-theorems}, @@ -3334,9 +3331,9 @@ \label{ssec:size}\ text \ -For each datatype @{text t}, the \hthm{size} plugin generates a generic size -function @{text "t.size_t"} as well as a specific instance -@{text "size :: t \ nat"} belonging to the @{text size} type class. The +For each datatype \t\, the \hthm{size} plugin generates a generic size +function \t.size_t\ as well as a specific instance +\size :: t \ nat\ belonging to the \size\ type class. The \keyw{fun} command relies on @{const size} to prove termination of recursive functions on datatypes. @@ -3345,20 +3342,20 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{size} @{text "[simp, code]"}\rm:] ~ \\ +\item[\t.\\hthm{size} \[simp, code]\\rm:] ~ \\ @{thm list.size(1)[no_vars]} \\ @{thm list.size(2)[no_vars]} \\ @{thm list.size(3)[no_vars]} \\ @{thm list.size(4)[no_vars]} -\item[@{text "t."}\hthm{size_gen}\rm:] ~ \\ +\item[\t.\\hthm{size_gen}\rm:] ~ \\ @{thm list.size_gen(1)[no_vars]} \\ @{thm list.size_gen(2)[no_vars]} -\item[@{text "t."}\hthm{size_gen_o_map}\rm:] ~ \\ +\item[\t.\\hthm{size_gen_o_map}\rm:] ~ \\ @{thm list.size_gen_o_map[no_vars]} -\item[@{text "t."}\hthm{size_neq}\rm:] ~ \\ +\item[\t.\\hthm{size_neq}\rm:] ~ \\ This property is missing for @{typ "'a list"}. If the @{term size} function always evaluates to a non-zero value, this theorem has the form @{prop "\ size x = 0"}. @@ -3366,14 +3363,14 @@ \end{description} \end{indentblock} -The @{text "t.size"} and @{text "t.size_t"} functions generated for datatypes -defined by nested recursion through a datatype @{text u} depend on -@{text "u.size_u"}. - -If the recursion is through a non-datatype @{text u} with type arguments -@{text "'a\<^sub>1, \, 'a\<^sub>m"}, by default @{text u} values are given a size of 0. This +The \t.size\ and \t.size_t\ functions generated for datatypes +defined by nested recursion through a datatype \u\ depend on +\u.size_u\. + +If the recursion is through a non-datatype \u\ with type arguments +\'a\<^sub>1, \, 'a\<^sub>m\, by default \u\ values are given a size of 0. This can be improved upon by registering a custom size function of type -@{text "('a\<^sub>1 \ nat) \ \ \ ('a\<^sub>m \ nat) \ u \ nat"} using +\('a\<^sub>1 \ nat) \ \ \ ('a\<^sub>m \ nat) \ u \ nat\ using the ML function @{ML BNF_LFP_Size.register_size} or @{ML BNF_LFP_Size.register_size_global}. See theory \<^file>\~~/src/HOL/Library/Multiset.thy\ for an example. @@ -3385,7 +3382,7 @@ text \ For each (co)datatype with live type arguments and each manually registered BNF, -the \hthm{transfer} plugin generates a predicator @{text "t.pred_t"} and +the \hthm{transfer} plugin generates a predicator \t.pred_t\ and properties that guide the Transfer tool. For types with at least one live type argument and \emph{no dead type @@ -3394,48 +3391,48 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{Domainp_rel} @{text "[relator_domain]"}\rm:] ~ \\ +\item[\t.\\hthm{Domainp_rel} \[relator_domain]\\rm:] ~ \\ @{thm list.Domainp_rel[no_vars]} -\item[@{text "t."}\hthm{left_total_rel} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{left_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.left_total_rel[no_vars]} -\item[@{text "t."}\hthm{left_unique_rel} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{left_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.left_unique_rel[no_vars]} -\item[@{text "t."}\hthm{right_total_rel} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{right_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.right_total_rel[no_vars]} -\item[@{text "t."}\hthm{right_unique_rel} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{right_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.right_unique_rel[no_vars]} -\item[@{text "t."}\hthm{bi_total_rel} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{bi_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.bi_total_rel[no_vars]} -\item[@{text "t."}\hthm{bi_unique_rel} @{text "[transfer_rule]"}\rm:] ~ \\ +\item[\t.\\hthm{bi_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.bi_unique_rel[no_vars]} \end{description} \end{indentblock} For (co)datatypes with at least one live type argument, the plugin sets the -@{text "[transfer_rule]"} attribute on the following (co)datatypes properties: -@{text "t.case_"}\allowbreak @{text "transfer"}, -@{text "t.sel_"}\allowbreak @{text "transfer"}, -@{text "t.ctr_"}\allowbreak @{text "transfer"}, -@{text "t.disc_"}\allowbreak @{text "transfer"}, -@{text "t.rec_"}\allowbreak @{text "transfer"}, and -@{text "t.corec_"}\allowbreak @{text "transfer"}. +\[transfer_rule]\ attribute on the following (co)datatypes properties: +\t.case_\\allowbreak \transfer\, +\t.sel_\\allowbreak \transfer\, +\t.ctr_\\allowbreak \transfer\, +\t.disc_\\allowbreak \transfer\, +\t.rec_\\allowbreak \transfer\, and +\t.corec_\\allowbreak \transfer\. For (co)datatypes that further have \emph{no dead type arguments}, the plugin -sets @{text "[transfer_rule]"} on -@{text "t.set_"}\allowbreak @{text "transfer"}, -@{text "t.map_"}\allowbreak @{text "transfer"}, and -@{text "t.rel_"}\allowbreak @{text "transfer"}. +sets \[transfer_rule]\ on +\t.set_\\allowbreak \transfer\, +\t.map_\\allowbreak \transfer\, and +\t.rel_\\allowbreak \transfer\. For @{command primrec}, @{command primcorec}, and @{command primcorecursive}, -the plugin implements the generation of the @{text "f.transfer"} property, -conditioned by the @{text transfer} option, and sets the -@{text "[transfer_rule]"} attribute on these. +the plugin implements the generation of the \f.transfer\ property, +conditioned by the \transfer\ option, and sets the +\[transfer_rule]\ attribute on these. \ @@ -3452,16 +3449,16 @@ \begin{indentblock} \begin{description} -\item[@{text "t."}\hthm{Quotient} @{text "[quot_map]"}\rm:] ~ \\ +\item[\t.\\hthm{Quotient} \[quot_map]\\rm:] ~ \\ @{thm list.Quotient[no_vars]} \end{description} \end{indentblock} -In addition, the plugin sets the @{text "[relator_eq]"} attribute on a -variant of the @{text t.rel_eq_onp} property, the @{text "[relator_mono]"} -attribute on @{text t.rel_mono}, and the @{text "[relator_distr]"} attribute -on @{text t.rel_compp}. +In addition, the plugin sets the \[relator_eq]\ attribute on a +variant of the \t.rel_eq_onp\ property, the \[relator_mono]\ +attribute on \t.rel_mono\, and the \[relator_distr]\ attribute +on \t.rel_compp\. \ @@ -3522,12 +3519,12 @@ \item \emph{The \emph{\keyw{primcorec}} command does not allow corecursion under -@{text "case"}--@{text "of"} for datatypes that are defined without +\case\--\of\ for datatypes that are defined without discriminators and selectors.} \item \emph{There is no way to use an overloaded constant from a syntactic type -class, such as @{text 0}, as a constructor.} +class, such as \0\, as a constructor.} \item \emph{There is no way to register the same type as both a datatype and a @@ -3553,7 +3550,7 @@ (co)datatype package. Andreas Lochbihler provided lots of comments on earlier versions of the package, especially on the coinductive part. Brian Huffman suggested major simplifications to the internal constructions. Ond\v{r}ej -Kun\v{c}ar implemented the @{text transfer} and @{text lifting} plugins. +Kun\v{c}ar implemented the \transfer\ and \lifting\ plugins. Christian Sternagel and Ren\'e Thiemann ported the \keyw{derive} command from the \emph{Archive of Formal Proofs} to the new datatypes. Gerwin Klein and Lars Noschinski implemented the @{command simps_of_case} and @{command