diff --git a/src/Doc/Datatypes/Datatypes.thy b/src/Doc/Datatypes/Datatypes.thy --- a/src/Doc/Datatypes/Datatypes.thy +++ b/src/Doc/Datatypes/Datatypes.thy @@ -1,3679 +1,3681 @@ (* Title: Doc/Datatypes/Datatypes.thy Author: Julian Biendarra, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, Ecole de technologie superieure Author: Lorenz Panny, TU Muenchen Author: Andrei Popescu, TU Muenchen Author: Dmitriy Traytel, TU Muenchen Tutorial for (co)datatype definitions. *) theory Datatypes imports Setup "HOL-Library.BNF_Axiomatization" "HOL-Library.Countable" "HOL-Library.FSet" "HOL-Library.Simps_Case_Conv" begin (*<*) unbundle cardinal_syntax (*>*) section \Introduction \label{sec:introduction}\ text \ The 2013 edition of Isabelle introduced a definitional package for freely generated datatypes and codatatypes. This package replaces the earlier implementation due to Berghofer and Wenzel @{cite "Berghofer-Wenzel:1999:TPHOL"}. Perhaps the main advantage of the new package is that it supports recursion through a large class of non-datatypes, such as finite sets: \ 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") text \ \noindent Another strong point is the support for local definitions: \ context linorder begin datatype flag = Less | Eq | Greater end text \ \noindent Furthermore, the package provides a lot of convenience, including automatically generated discriminators, selectors, and relators as well as a wealth of properties about them. In addition to inductive datatypes, the package supports coinductive datatypes, or \emph{codatatypes}, which allow infinite values. For example, the following command introduces the type of lazy lists, which comprises both finite and infinite values: \ (*<*) locale early locale late (*>*) codatatype (*<*)(in early) (*>*)'a llist = LNil | LCons 'a "'a llist" text \ \noindent Mixed inductive--coinductive recursion is possible via nesting. Compare the following four Rose tree examples: \ datatype (*<*)(in early) (*>*)'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f 'a "'a tree\<^sub>f\<^sub>f list" datatype (*<*)(in early) (*>*)'a tree\<^sub>f\<^sub>i = Node\<^sub>f\<^sub>i 'a "'a tree\<^sub>f\<^sub>i llist" codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>f = Node\<^sub>i\<^sub>f 'a "'a tree\<^sub>i\<^sub>f list" codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist" text \ The first two tree types allow only paths of finite length, whereas the last two allow infinite paths. Orthogonally, the nodes in the first and third types have finitely many direct subtrees, whereas those of the second and fourth may have infinite branching. The package is part of \<^theory>\Main\. Additional functionality is provided by the theory \<^file>\~~/src/HOL/Library/BNF_Axiomatization.thy\. The package, like its predecessor, fully adheres to the LCF philosophy @{cite mgordon79}: The characteristic theorems associated with the specified (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 \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"}. The central notion is that of a \emph{bounded natural functor} (BNF)---a well-behaved type constructor for which nested (co)recursion is supported. This tutorial is organized as follows: \begin{itemize} \setlength{\itemsep}{0pt} \item Section \ref{sec:defining-datatypes}, ``Defining Datatypes,'' describes how to specify datatypes using the @{command datatype} command. \item Section \ref{sec:defining-primitively-recursive-functions}, ``Defining Primitively Recursive Functions,'' describes how to specify functions using @{command primrec}. (A separate tutorial @{cite "isabelle-function"} describes the more powerful \keyw{fun} and \keyw{function} commands.) \item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,'' describes how to specify codatatypes using the @{command codatatype} command. \item Section \ref{sec:defining-primitively-corecursive-functions}, ``Defining Primitively Corecursive Functions,'' describes how to specify functions using the @{command primcorec} and @{command primcorecursive} commands. (A separate tutorial @{cite "isabelle-corec"} describes the more powerful \keyw{corec} and \keyw{corecursive} commands.) \item Section \ref{sec:registering-bounded-natural-functors}, ``Registering Bounded Natural Functors,'' explains how to use the @{command bnf} command to register arbitrary type constructors as BNFs. \item Section \ref{sec:deriving-destructors-and-theorems-for-free-constructors}, ``Deriving Destructors and Theorems for Free Constructors,'' explains how to use the command @{command free_constructors} to derive destructor constants and theorems for freely generated types, as performed internally by @{command datatype} and @{command codatatype}. %\item Section \ref{sec:using-the-standard-ml-interface}, ``Using the Standard %ML Interface,'' describes the package's programmatic interface. \item Section \ref{sec:selecting-plugins}, ``Selecting Plugins,'' is concerned with the package's interoperability with other Isabelle packages and tools, such as the code generator, Transfer, Lifting, and Quickcheck. \item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and Limitations,'' concludes with known open issues. \end{itemize} \newbox\boxA \setbox\boxA=\hbox{\texttt{NOSPAM}} \newcommand\authoremaili{\texttt{jasmin.blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak gmail.\allowbreak com}} Comments and bug reports concerning either the package or this tutorial should be directed to the second author at \authoremaili{} or to the \texttt{cl-isabelle-users} mailing list. \ section \Defining Datatypes \label{sec:defining-datatypes}\ text \ Datatypes can be specified using the @{command datatype} command. \ subsection \Introductory Examples \label{ssec:datatype-introductory-examples}\ text \ Datatypes are illustrated through concrete examples featuring different flavors of recursion. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. \ subsubsection \Nonrecursive Types \label{sssec:datatype-nonrecursive-types}\ text \ Datatypes are introduced by specifying the desired names and argument types for their constructors. \emph{Enumeration} types are the simplest form of datatype. All their constructors are nullary: \ datatype trool = Truue | Faalse | Perhaaps text \ \noindent \<^const>\Truue\, \<^const>\Faalse\, and \<^const>\Perhaaps\ have the type \<^typ>\trool\. Polymorphic types are possible, such as the following option type, modeled after its homologue from the \<^theory>\HOL.Option\ theory: \ (*<*) hide_const None Some map_option hide_type option (*>*) datatype 'a option = None | Some 'a text \ \noindent The constructors are \None :: 'a option\ and \Some :: 'a \ 'a option\. The next example has three type parameters: \ datatype ('a, 'b, 'c) triple = Triple 'a 'b 'c text \ \noindent The constructor is \Triple :: 'a \ 'b \ 'c \ ('a, 'b, 'c) triple\. Unlike in Standard ML, curried constructors are supported. The uncurried variant is also possible: \ datatype ('a, 'b, 'c) triple\<^sub>u = Triple\<^sub>u "'a * 'b * 'c" text \ \noindent Occurrences of nonatomic types on the right-hand side of the equal sign must be enclosed in double quotes, as is customary in Isabelle. \ subsubsection \Simple Recursion \label{sssec:datatype-simple-recursion}\ text \ Natural numbers are the simplest example of a recursive type: \ datatype nat = Zero | Succ nat text \ \noindent Lists were shown in the introduction. Terminated lists are a variant that stores a value of type \<^typ>\'b\ at the very end: \ datatype (*<*)(in early) (*>*)('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist" subsubsection \Mutual Recursion \label{sssec:datatype-mutual-recursion}\ text \ \emph{Mutually recursive} types are introduced simultaneously and may refer to each other. The example below introduces a pair of types for even and odd natural numbers: \ datatype even_nat = Even_Zero | Even_Succ odd_nat and odd_nat = Odd_Succ even_nat text \ \noindent Arithmetic expressions are defined via terms, terms via factors, and factors via expressions: \ datatype ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp" and ('a, 'b) trm = Factor "('a, 'b) fct" | Prod "('a, 'b) fct" "('a, 'b) trm" and ('a, 'b) fct = Const 'a | Var 'b | Expr "('a, 'b) exp" subsubsection \Nested Recursion \label{sssec:datatype-nested-recursion}\ text \ \emph{Nested recursion} occurs when recursive occurrences of a type appear under a type constructor. The introduction showed some examples of trees with nesting through lists. A more complex example, that reuses our \<^type>\option\ type, follows: \ datatype 'a btree = BNode 'a "'a btree option" "'a btree option" text \ \noindent Not all nestings are admissible. For example, this command will fail: \ datatype 'a wrong = W1 | W2 (*<*)'a typ (*>*)"'a wrong \ 'a" text \ \noindent 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~\\\: \ datatype ('a, 'b) fun_copy = Fun "'a \ 'b" datatype 'a also_wrong = W1 | W2 (*<*)'a typ (*>*)"('a also_wrong, 'a) fun_copy" text \ \noindent The following definition of \<^typ>\'a\-branching trees is legal: \ datatype 'a ftree = FTLeaf 'a | FTNode "'a \ 'a ftree" text \ \noindent And so is the definition of hereditarily finite sets: \ datatype hfset = HFSet "hfset fset" text \ \noindent 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. Type constructors must be registered as BNFs to have live arguments. This is done automatically for datatypes and codatatypes introduced by the @{command datatype} and @{command codatatype} commands. Section~\ref{sec:registering-bounded-natural-functors} explains how to register arbitrary type constructors as BNFs. Here is another example that fails: \ datatype 'a pow_list = PNil 'a (*<*)'a datatype 'a pow_list' = PNil' 'a (*>*)| PCons "('a * 'a) pow_list" text \ \noindent This attempted definition features a different flavor of nesting, where the recursive call in the type specification occurs around (rather than inside) another type constructor. \ subsubsection \Auxiliary Constants \label{sssec:datatype-auxiliary-constants}\ text \ 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 \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\: \ (*<*) no_translations "[x, xs]" == "x # [xs]" "[x]" == "x # []" no_notation Nil ("[]") and Cons (infixr "#" 65) hide_type list hide_const Nil Cons case_list hd tl set map list_all2 rec_list size_list list_all context early begin (*>*) datatype (set: 'a) list = null: Nil | Cons (hd: 'a) (tl: "'a list") for map: map rel: list_all2 pred: list_all where "tl Nil = Nil" text \ \noindent The types of the constants that appear in the specification are listed below. \medskip \begin{tabular}{@ {}ll@ {}} Constructors: & \Nil :: 'a list\ \\ & \Cons :: 'a \ 'a list \ 'a list\ \\ Discriminator: & \null :: 'a list \ bool\ \\ Selectors: & \hd :: 'a list \ 'a\ \\ & \tl :: 'a list \ 'a list\ \\ Set function: & \set :: 'a list \ 'a set\ \\ Map function: & \map :: ('a \ 'b) \ 'a list \ 'b list\ \\ Relator: & \list_all2 :: ('a \ 'b \ bool) \ 'a list \ 'b list \ bool\ \end{tabular} \medskip The discriminator \<^const>\null\ and the selectors \<^const>\hd\ and \<^const>\tl\ are characterized by the following conditional equations: % \[@{thm list.collapse(1)[of xs, no_vars]} \qquad @{thm list.collapse(2)[of xs, no_vars]}\] % For two-constructor datatypes, a single discriminator constant is sufficient. The discriminator associated with \<^const>\Cons\ is simply \<^term>\\xs. \ null xs\. The \keyw{where} clause at the end of the command specifies a default value for selectors applied to constructors on which they are not a priori specified. In the example, it is used to ensure that the tail of the empty list is itself (instead of being left unspecified). Because \<^const>\Nil\ is nullary, it is also possible to use \<^term>\\xs. xs = Nil\ as a discriminator. This is the default behavior if we omit the identifier \<^const>\null\ and the associated colon. Some users argue against this, because the mixture of constructors and selectors in the characteristic theorems can lead Isabelle's automation to switch between the constructor and the destructor view in surprising ways. The usual mixfix syntax annotations are available for both types and constructors. For example: \ (*<*) end (*>*) datatype ('a, 'b) prod (infixr "*" 20) = Pair 'a 'b text \\blankline\ datatype (set: 'a) list = null: Nil ("[]") | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) for map: map rel: list_all2 pred: list_all text \ \noindent Incidentally, this is how the traditional syntax can be set up: \ syntax "_list" :: "args \ 'a list" ("[(_)]") text \\blankline\ translations "[x, xs]" == "x # [xs]" "[x]" == "x # []" subsection \Command Syntax \label{ssec:datatype-command-syntax}\ subsubsection \\keyw{datatype} \label{sssec:datatype}\ text \ \begin{matharray}{rcl} @{command_def "datatype"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command datatype} target? @{syntax dt_options}? @{syntax dt_spec} ; @{syntax_def dt_options}: '(' ((@{syntax plugins} | 'discs_sels') + ',') ')' ; @{syntax_def plugins}: 'plugins' ('only' | 'del') ':' (name +) ; @{syntax_def dt_spec}: (@{syntax dt_name} '=' (@{syntax dt_ctor} + '|') \ @{syntax map_rel_pred}? (@'where' (prop + '|'))? + @'and') ; @{syntax_def map_rel_pred}: @'for' ((('map' | 'rel' | 'pred') ':' name) +) \ \medskip \noindent The @{command datatype} command introduces a set of mutually recursive datatypes specified by their constructors. The syntactic entity \synt{target} can be used to specify a local 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 options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item 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 \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: \<^rail>\ @{syntax_def dt_name}: @{syntax tyargs}? name mixfix? ; @{syntax_def tyargs}: typefree | '(' (('dead' | name ':')? typefree + ',') ')' \ \medskip \noindent The syntactic entity \synt{name} denotes an identifier, \synt{mixfix} denotes the usual parenthesized mixfix notation, and \synt{typefree} denotes fixed type 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 (\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 type argument. Inside a mutually recursive specification, all defined datatypes must mention exactly the same type variables in the same order. \<^rail>\ @{syntax_def dt_ctor}: (name ':')? name (@{syntax dt_ctor_arg} * ) mixfix? \ \medskip \noindent 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 \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 ')' \ \medskip \noindent The syntactic entity \synt{type} denotes a HOL type @{cite "isabelle-isar-ref"}. 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 \discs_sels\ option) but no name is supplied, the default name is \un_C\<^sub>ji\. \ subsubsection \\keyw{datatype_compat} \label{sssec:datatype-compat}\ text \ \begin{matharray}{rcl} @{command_def "datatype_compat"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command datatype_compat} (name +) \ \medskip \noindent The @{command datatype_compat} command registers new-style datatypes as old-style datatypes and invokes the old-style plugins. For example: \ datatype_compat even_nat odd_nat text \\blankline\ ML \Old_Datatype_Data.get_info \<^theory> \<^type_name>\even_nat\\ text \ The syntactic entity \synt{name} denotes an identifier @{cite "isabelle-isar-ref"}. The command is sometimes useful when migrating from the old datatype package to the new one. A few remarks concern nested recursive datatypes: \begin{itemize} \setlength{\itemsep}{0pt} \item The old-style, nested-as-mutual induction rule and recursor theorems are 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 \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. \end{itemize} \ subsection \Generated Constants \label{ssec:datatype-generated-constants}\ text \ 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: & \t.case_t\ (rendered using the familiar \case\--\of\ syntax) \\ Discriminators: & \t.is_C\<^sub>1\$, \ldots, $\t.is_C\<^sub>n\ \\ Selectors: & \t.un_C\<^sub>11\$, \ldots, $\t.un_C\<^sub>1k\<^sub>1\ \\ & \quad\vdots \\ & \t.un_C\<^sub>n1\$, \ldots, $\t.un_C\<^sub>nk\<^sub>n\ \\ Set functions: & \t.set\<^sub>1_t\, \ldots, \t.set\<^sub>m_t\ \\ Map function: & \t.map_t\ \\ Relator: & \t.rel_t\ \\ Recursor: & \t.rec_t\ \end{tabular} \medskip \noindent 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 ``\t.\'' is an optional component of the names and is normally hidden. \ subsection \Generated Theorems \label{ssec:datatype-generated-theorems}\ text \ The characteristic theorems generated by @{command datatype} are grouped in three broad categories: \begin{itemize} \setlength{\itemsep}{0pt} \item The \emph{free constructor theorems} (Section~\ref{sssec:free-constructor-theorems}) are properties of the constructors and destructors that can be derived for any freely generated type. Internally, the derivation is performed by @{command free_constructors}. \item The \emph{functorial theorems} (Section~\ref{sssec:functorial-theorems}) are properties of datatypes related to their BNF nature. \item The \emph{inductive theorems} (Section~\ref{sssec:inductive-theorems}) are properties of datatypes related to their inductive nature. \end{itemize} \noindent The full list of named theorems can be obtained by issuing the command \keyw{print_theorems} immediately after the datatype definition. This list includes theorems produced by plugins (Section~\ref{sec:selecting-plugins}), but normally excludes low-level theorems that reveal internal constructions. To make these accessible, add the line \ declare [[bnf_internals]] (*<*) declare [[bnf_internals = false]] (*>*) subsubsection \Free Constructor Theorems \label{sssec:free-constructor-theorems}\ (*<*) consts nonnull :: 'a (*>*) text \ The free constructor theorems are partitioned in three subgroups. The first subgroup of properties is concerned with the constructors. They are listed below for \<^typ>\'a list\: \begin{indentblock} \begin{description} \item[\t.\\hthm{inject} \[iff, induct_simp]\\rm:] ~ \\ @{thm list.inject[no_vars]} \item[\t.\\hthm{distinct} \[simp, induct_simp]\\rm:] ~ \\ @{thm list.distinct(1)[no_vars]} \\ @{thm list.distinct(2)[no_vars]} \item[\t.\\hthm{exhaust} \[cases t, case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust[no_vars]} \item[\t.\\hthm{nchotomy}\rm:] ~ \\ @{thm list.nchotomy[no_vars]} \end{description} \end{indentblock} \noindent In addition, these nameless theorems are registered as safe elimination rules: \begin{indentblock} \begin{description} \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]} \end{description} \end{indentblock} \noindent The next subgroup is concerned with the case combinator: \begin{indentblock} \begin{description} \item[\t.\\hthm{case} \[simp, code]\\rm:] ~ \\ @{thm list.case(1)[no_vars]} \\ @{thm list.case(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{case_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.case_cong[no_vars]} \item[\t.\\hthm{case_cong_weak} \[cong]\\rm:] ~ \\ @{thm list.case_cong_weak[no_vars]} \item[\t.\\hthm{case_distrib}\rm:] ~ \\ @{thm list.case_distrib[no_vars]} \item[\t.\\hthm{split}\rm:] ~ \\ @{thm list.split[no_vars]} \item[\t.\\hthm{split_asm}\rm:] ~ \\ @{thm list.split_asm[no_vars]} \item[\t.\\hthm{splits} = \split split_asm\] \end{description} \end{indentblock} \noindent The third subgroup revolves around discriminators and selectors: \begin{indentblock} \begin{description} \item[\t.\\hthm{disc} \[simp]\\rm:] ~ \\ @{thm list.disc(1)[no_vars]} \\ @{thm list.disc(2)[no_vars]} \item[\t.\\hthm{discI}\rm:] ~ \\ @{thm list.discI(1)[no_vars]} \\ @{thm list.discI(2)[no_vars]} \item[\t.\\hthm{sel} \[simp, code]\\rm:] ~ \\ @{thm list.sel(1)[no_vars]} \\ @{thm list.sel(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{collapse} \[simp]\\rm:] ~ \\ @{thm list.collapse(1)[no_vars]} \\ @{thm list.collapse(2)[no_vars]} \\ 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[\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[\t.\\hthm{exhaust_disc} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust_disc[no_vars]} \item[\t.\\hthm{exhaust_sel} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust_sel[no_vars]} \item[\t.\\hthm{expand}\rm:] ~ \\ @{thm list.expand[no_vars]} \item[\t.\\hthm{split_sel}\rm:] ~ \\ @{thm list.split_sel[no_vars]} \item[\t.\\hthm{split_sel_asm}\rm:] ~ \\ @{thm list.split_sel_asm[no_vars]} \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[\t.\\hthm{disc_eq_case}\rm:] ~ \\ @{thm list.disc_eq_case(1)[no_vars]} \\ @{thm list.disc_eq_case(2)[no_vars]} \end{description} \end{indentblock} \noindent 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}). \ subsubsection \Functorial Theorems \label{sssec:functorial-theorems}\ text \ The functorial theorems are generated for type constructors with at least one live type argument (e.g., \<^typ>\'a list\). They are partitioned in two subgroups. The first subgroup consists of properties involving the constructors or the destructors and either a set function, the map function, the predicator, or the relator: \begin{indentblock} \begin{description} \item[\t.\\hthm{case_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.case_transfer[no_vars]} \\ This property is generated by the \transfer\ plugin (Section~\ref{ssec:transfer}). %The \[transfer_rule]\ attribute is set by the \transfer\ plugin %(Section~\ref{ssec:transfer}). \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 \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}). \item[\t.\\hthm{ctr_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.ctr_transfer(1)[no_vars]} \\ @{thm list.ctr_transfer(2)[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) . \item[\t.\\hthm{disc_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.disc_transfer(1)[no_vars]} \\ @{thm list.disc_transfer(2)[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}). \item[\t.\\hthm{set} \[simp, code]\\rm:] ~ \\ @{thm list.set(1)[no_vars]} \\ @{thm list.set(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{set_cases} \[consumes 1, cases set: set\<^sub>i_t]\\rm:] ~ \\ @{thm list.set_cases[no_vars]} \item[\t.\\hthm{set_intros}\rm:] ~ \\ @{thm list.set_intros(1)[no_vars]} \\ @{thm list.set_intros(2)[no_vars]} \item[\t.\\hthm{set_sel}\rm:] ~ \\ @{thm list.set_sel[no_vars]} \item[\t.\\hthm{map} \[simp, code]\\rm:] ~ \\ @{thm list.map(1)[no_vars]} \\ @{thm list.map(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{map_disc_iff} \[simp]\\rm:] ~ \\ @{thm list.map_disc_iff[no_vars]} \item[\t.\\hthm{map_sel}\rm:] ~ \\ @{thm list.map_sel[no_vars]} \item[\t.\\hthm{pred_inject} \[simp]\\rm:] ~ \\ @{thm list.pred_inject(1)[no_vars]} \\ @{thm list.pred_inject(2)[no_vars]} \item[\t.\\hthm{rel_inject} \[simp]\\rm:] ~ \\ @{thm list.rel_inject(1)[no_vars]} \\ @{thm list.rel_inject(2)[no_vars]} \item[\t.\\hthm{rel_distinct} \[simp]\\rm:] ~ \\ @{thm list.rel_distinct(1)[no_vars]} \\ @{thm list.rel_distinct(2)[no_vars]} \item[\t.\\hthm{rel_intros}\rm:] ~ \\ @{thm list.rel_intros(1)[no_vars]} \\ @{thm list.rel_intros(2)[no_vars]} \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[\t.\\hthm{rel_sel}\rm:] ~ \\ @{thm list.rel_sel[no_vars]} \end{description} \end{indentblock} \noindent 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, the map function, the predicator, and the relator: \begin{indentblock} \begin{description} \item[\t.\\hthm{inj_map}\rm:] ~ \\ @{thm list.inj_map[no_vars]} \item[\t.\\hthm{inj_map_strong}\rm:] ~ \\ @{thm list.inj_map_strong[no_vars]} \item[\t.\\hthm{map_comp}\rm:] ~ \\ @{thm list.map_comp[no_vars]} \item[\t.\\hthm{map_cong0}\rm:] ~ \\ @{thm list.map_cong0[no_vars]} \item[\t.\\hthm{map_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.map_cong[no_vars]} \item[\t.\\hthm{map_cong_pred}\rm:] ~ \\ @{thm list.map_cong_pred[no_vars]} \item[\t.\\hthm{map_cong_simp}\rm:] ~ \\ @{thm list.map_cong_simp[no_vars]} \item[\t.\\hthm{map_id0}\rm:] ~ \\ @{thm list.map_id0[no_vars]} \item[\t.\\hthm{map_id}\rm:] ~ \\ @{thm list.map_id[no_vars]} \item[\t.\\hthm{map_ident}\rm:] ~ \\ @{thm list.map_ident[no_vars]} \item[\t.\\hthm{map_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.map_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \item[\t.\\hthm{pred_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.pred_cong[no_vars]} \item[\t.\\hthm{pred_cong_simp}\rm:] ~ \\ @{thm list.pred_cong_simp[no_vars]} \item[\t.\\hthm{pred_map}\rm:] ~ \\ @{thm list.pred_map[no_vars]} \item[\t.\\hthm{pred_mono_strong}\rm:] ~ \\ @{thm list.pred_mono_strong[no_vars]} \item[\t.\\hthm{pred_rel}\rm:] ~ \\ @{thm list.pred_rel[no_vars]} \item[\t.\\hthm{pred_set}\rm:] ~ \\ @{thm list.pred_set[no_vars]} \item[\t.\\hthm{pred_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.pred_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \item[\t.\\hthm{pred_True}\rm:] ~ \\ @{thm list.pred_True[no_vars]} \item[\t.\\hthm{set_map}\rm:] ~ \\ @{thm list.set_map[no_vars]} \item[\t.\\hthm{set_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.set_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \item[\t.\\hthm{rel_compp} \[relator_distr]\\rm:] ~ \\ @{thm list.rel_compp[no_vars]} \\ The \[relator_distr]\ attribute is set by the \lifting\ plugin (Section~\ref{ssec:lifting}). \item[\t.\\hthm{rel_conversep}\rm:] ~ \\ @{thm list.rel_conversep[no_vars]} \item[\t.\\hthm{rel_eq}\rm:] ~ \\ @{thm list.rel_eq[no_vars]} \item[\t.\\hthm{rel_eq_onp}\rm:] ~ \\ @{thm list.rel_eq_onp[no_vars]} \item[\t.\\hthm{rel_flip}\rm:] ~ \\ @{thm list.rel_flip[no_vars]} \item[\t.\\hthm{rel_map}\rm:] ~ \\ @{thm list.rel_map(1)[no_vars]} \\ @{thm list.rel_map(2)[no_vars]} \item[\t.\\hthm{rel_mono} \[mono, relator_mono]\\rm:] ~ \\ @{thm list.rel_mono[no_vars]} \\ The \[relator_mono]\ attribute is set by the \lifting\ plugin (Section~\ref{ssec:lifting}). \item[\t.\\hthm{rel_mono_strong}\rm:] ~ \\ @{thm list.rel_mono_strong[no_vars]} \item[\t.\\hthm{rel_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.rel_cong[no_vars]} \item[\t.\\hthm{rel_cong_simp}\rm:] ~ \\ @{thm list.rel_cong_simp[no_vars]} \item[\t.\\hthm{rel_refl}\rm:] ~ \\ @{thm list.rel_refl[no_vars]} \item[\t.\\hthm{rel_refl_strong}\rm:] ~ \\ @{thm list.rel_refl_strong[no_vars]} \item[\t.\\hthm{rel_reflp}\rm:] ~ \\ @{thm list.rel_reflp[no_vars]} \item[\t.\\hthm{rel_symp}\rm:] ~ \\ @{thm list.rel_symp[no_vars]} \item[\t.\\hthm{rel_transp}\rm:] ~ \\ @{thm list.rel_transp[no_vars]} \item[\t.\\hthm{rel_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.rel_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} \end{indentblock} \ subsubsection \Inductive Theorems \label{sssec:inductive-theorems}\ text \ The inductive theorems are as follows: \begin{indentblock} \begin{description} \item[\t.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n, induct t]\\rm:] ~ \\ @{thm list.induct[no_vars]} \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@ {}} \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[\t.\\hthm{rec} \[simp, code]\\rm:] ~ \\ @{thm list.rec(1)[no_vars]} \\ @{thm list.rec(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{rec_o_map}\rm:] ~ \\ @{thm list.rec_o_map[no_vars]} \item[\t.\\hthm{rec_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.rec_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} \end{indentblock} \noindent For convenience, @{command datatype} also provides the following collection: \begin{indentblock} \begin{description} \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} \ subsection \Proof Method \label{ssec:datatype-proof-method}\ subsubsection \\textit{countable_datatype} \label{sssec:datatype-compat}\ text \ The theory \<^file>\~~/src/HOL/Library/Countable.thy\ provides a proof method called @{method countable_datatype} that can be used to prove the countability of many datatypes, building on the countability of the types appearing in their definitions and of any type arguments. For example: \ instance list :: (countable) countable by countable_datatype subsection \Antiquotation \label{ssec:datatype-antiquotation}\ subsubsection \\textit{datatype} \label{sssec:datatype-datatype}\ text \ The \textit{datatype} antiquotation, written \texttt{\char`\\\char`<\char`^}\verb|datatype>|\guilsinglleft\textit{t}\guilsinglright{} or \texttt{@}\verb|{datatype| \textit{t}\verb|}|, where \textit{t} is a type name, expands to \LaTeX{} code for the definition of the datatype, with each constructor listed with its argument types. For example, if \textit{t} is @{type option}: \begin{quote} \<^datatype>\option\ \end{quote} \ subsection \Compatibility Issues \label{ssec:datatype-compatibility-issues}\ text \ The command @{command datatype} has been designed to be highly compatible with the old, pre-Isabelle2015 command, to ease migration. There are nonetheless a few incompatibilities that may arise when porting: \begin{itemize} \setlength{\itemsep}{0pt} \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 \BNF_LFP_Compat\ structure provides convenience functions that simulate the old interfaces in terms of the new ones. \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 handled specially. In the new package, nested recursion (for functions and non-functions) is handled in a more modular fashion. The old-style recursor can be generated on demand 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}. \item \emph{Accordingly, the induction rule is different for nested recursive datatypes.} Again, the old-style induction rule can be generated on demand 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 \[unfolded all_mem_range]\ attribute on \t.induct\. \item \emph{The \<^const>\size\ function has a slightly different definition.} 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 \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 \size\ plugin (Section~\ref{sec:selecting-plugins}) and instantiating the \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 be difficult to port. \item \emph{Some constants and theorems have different names.} For non-mutually recursive datatypes, the alias \t.inducts\ for \t.induct\ is no longer generated. For $m > 1$ mutually recursive datatypes, \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 \[where \]\ attribute. The solution is to use the more robust \[of \]\ syntax. \end{itemize} \ section \Defining Primitively Recursive Functions \label{sec:defining-primitively-recursive-functions}\ text \ Recursive functions over datatypes can be specified using the @{command primrec} command, which supports primitive recursion, or using the \keyw{fun}, \keyw{function}, and \keyw{partial_function} commands. In this tutorial, the focus is on @{command primrec}; \keyw{fun} and \keyw{function} are described in a separate tutorial @{cite "isabelle-function"}. Because it is restricted to primitive recursion, @{command primrec} is less powerful than \keyw{fun} and \keyw{function}. However, there are primitively recursive specifications (e.g., based on infinitely branching or mutually recursive datatypes) for which \keyw{fun}'s termination check fails. It is also good style to use the simpler @{command primrec} mechanism when it works, both as an optimization and as documentation. \ subsection \Introductory Examples \label{ssec:primrec-introductory-examples}\ text \ Primitive recursion is illustrated through concrete examples based on the datatypes defined in Section~\ref{ssec:datatype-introductory-examples}. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. \ subsubsection \Nonrecursive Types \label{sssec:primrec-nonrecursive-types}\ text \ Primitive recursion removes one layer of constructors on the left-hand side in each equation. For example: \ primrec (nonexhaustive) bool_of_trool :: "trool \ bool" where "bool_of_trool Faalse \ False" | "bool_of_trool Truue \ True" text \\blankline\ primrec the_list :: "'a option \ 'a list" where "the_list None = []" | "the_list (Some a) = [a]" text \\blankline\ primrec the_default :: "'a \ 'a option \ 'a" where "the_default d None = d" | "the_default _ (Some a) = a" text \\blankline\ primrec mirrror :: "('a, 'b, 'c) triple \ ('c, 'b, 'a) triple" where "mirrror (Triple a b c) = Triple c b a" text \ \noindent The equations can be specified in any order, and it is acceptable to leave out some cases, which are then unspecified. Pattern matching on the left-hand side is restricted to a single datatype, which must correspond to the same argument in all equations. \ subsubsection \Simple Recursion \label{sssec:primrec-simple-recursion}\ text \ For simple recursive types, recursive calls on a constructor argument are allowed on the right-hand side: \ primrec replicate :: "nat \ 'a \ 'a list" where "replicate Zero _ = []" | "replicate (Succ n) x = x # replicate n x" text \\blankline\ primrec (nonexhaustive) at :: "'a list \ nat \ 'a" where "at (x # xs) j = (case j of Zero \ x | Succ j' \ at xs j')" text \\blankline\ primrec (*<*)(in early) (transfer) (*>*)tfold :: "('a \ 'b \ 'b) \ ('a, 'b) tlist \ 'b" where "tfold _ (TNil y) = y" | "tfold f (TCons x xs) = f x (tfold f xs)" text \ \noindent Pattern matching is only available for the argument on which the recursion takes place. Fortunately, it is easy to generate pattern-maching equations using the @{command simps_of_case} command provided by the theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\. \ simps_of_case at_simps_alt: at.simps text \ This generates the lemma collection @{thm [source] at_simps_alt}: % \[@{thm at_simps_alt(1)[no_vars]} \qquad @{thm at_simps_alt(2)[no_vars]}\] % The next example is defined using \keyw{fun} to escape the syntactic restrictions imposed on primitively recursive functions: \ fun at_least_two :: "nat \ bool" where "at_least_two (Succ (Succ _)) \ True" | "at_least_two _ \ False" subsubsection \Mutual Recursion \label{sssec:primrec-mutual-recursion}\ text \ The syntax for mutually recursive functions over mutually recursive datatypes is straightforward: \ primrec nat_of_even_nat :: "even_nat \ nat" and nat_of_odd_nat :: "odd_nat \ nat" where "nat_of_even_nat Even_Zero = Zero" | "nat_of_even_nat (Even_Succ n) = Succ (nat_of_odd_nat n)" | "nat_of_odd_nat (Odd_Succ n) = Succ (nat_of_even_nat n)" text \\blankline\ primrec eval\<^sub>e :: "('a \ int) \ ('b \ int) \ ('a, 'b) exp \ int" and eval\<^sub>t :: "('a \ int) \ ('b \ int) \ ('a, 'b) trm \ int" and eval\<^sub>f :: "('a \ int) \ ('b \ int) \ ('a, 'b) fct \ int" where "eval\<^sub>e \ \ (Term t) = eval\<^sub>t \ \ t" | "eval\<^sub>e \ \ (Sum t e) = eval\<^sub>t \ \ t + eval\<^sub>e \ \ e" | "eval\<^sub>t \ \ (Factor f) = eval\<^sub>f \ \ f" | "eval\<^sub>t \ \ (Prod f t) = eval\<^sub>f \ \ f + eval\<^sub>t \ \ t" | "eval\<^sub>f \ _ (Const a) = \ a" | "eval\<^sub>f _ \ (Var b) = \ b" | "eval\<^sub>f \ \ (Expr e) = eval\<^sub>e \ \ e" text \ \noindent Mutual recursion is possible within a single type, using \keyw{fun}: \ fun even :: "nat \ bool" and odd :: "nat \ bool" where "even Zero = True" | "even (Succ n) = odd n" | "odd Zero = False" | "odd (Succ n) = even n" subsubsection \Nested Recursion \label{sssec:primrec-nested-recursion}\ text \ In a departure from the old datatype package, nested recursion is normally handled via the map functions of the nesting type constructors. For example, recursive calls are lifted to lists using \<^const>\map\: \ (*<*) 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") (*>*) primrec at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \ nat list \ 'a" where "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js = (case js of [] \ a | j # js' \ at (map (\t. at\<^sub>f\<^sub>f t js') ts) j)" text \ \noindent 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\: \ primrec (*<*)(in early) (*>*)sum_btree :: "('a::{zero,plus}) btree \ 'a" where "sum_btree (BNode a lt rt) = a + the_default 0 (map_option sum_btree lt) + the_default 0 (map_option sum_btree rt)" text \ \noindent The same principle applies for arbitrary type constructors through which recursion is possible. Notably, the map function for the function type (\\\) is simply composition (\(\)\): \ primrec (*<*)(in early) (*>*)relabel_ft :: "('a \ 'a) \ 'a ftree \ 'a ftree" where "relabel_ft f (FTLeaf x) = FTLeaf (f x)" | "relabel_ft f (FTNode g) = FTNode (relabel_ft f \ g)" text \ \noindent For convenience, recursion through functions can also be expressed using $\lambda$-abstractions and function application rather than through composition. For example: \ primrec relabel_ft :: "('a \ 'a) \ 'a ftree \ 'a ftree" where "relabel_ft f (FTLeaf x) = FTLeaf (f x)" | "relabel_ft f (FTNode g) = FTNode (\x. relabel_ft f (g x))" text \\blankline\ primrec (nonexhaustive) subtree_ft :: "'a \ 'a ftree \ 'a ftree" where "subtree_ft x (FTNode g) = g x" text \ \noindent For recursion through curried $n$-ary functions, $n$ applications of \<^term>\(\)\ are necessary. The examples below illustrate the case where $n = 2$: \ datatype 'a ftree2 = FTLeaf2 'a | FTNode2 "'a \ 'a \ 'a ftree2" text \\blankline\ primrec (*<*)(in early) (*>*)relabel_ft2 :: "('a \ 'a) \ 'a ftree2 \ 'a ftree2" where "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" | "relabel_ft2 f (FTNode2 g) = FTNode2 ((\) ((\) (relabel_ft2 f)) g)" text \\blankline\ primrec relabel_ft2 :: "('a \ 'a) \ 'a ftree2 \ 'a ftree2" where "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" | "relabel_ft2 f (FTNode2 g) = FTNode2 (\x y. relabel_ft2 f (g x y))" text \\blankline\ primrec (nonexhaustive) subtree_ft2 :: "'a \ 'a \ 'a ftree2 \ 'a ftree2" where "subtree_ft2 x y (FTNode2 g) = g x y" text \ For any datatype featuring nesting, the predicator can be used instead of the map function, typically when defining predicates. For example: \ primrec increasing_tree :: "int \ int tree\<^sub>f\<^sub>f \ bool" where "increasing_tree m (Node\<^sub>f\<^sub>f n ts) \ n \ m \ list_all (increasing_tree (n + 1)) ts" subsubsection \Nested-as-Mutual Recursion \label{sssec:primrec-nested-as-mutual-recursion}\ (*<*) locale n2m begin (*>*) text \ For compatibility with the old package, but also because it is sometimes convenient in its own right, it is possible to treat nested recursive datatypes as mutually recursive ones if the recursion takes place though new-style datatypes. For example: \ primrec (nonexhaustive) at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \ nat list \ 'a" and ats\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f list \ nat \ nat list \ 'a" where "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js = (case js of [] \ a | j # js' \ ats\<^sub>f\<^sub>f ts j js')" | "ats\<^sub>f\<^sub>f (t # ts) j = (case j of Zero \ at\<^sub>f\<^sub>f t | Succ j' \ ats\<^sub>f\<^sub>f ts j')" text \ \noindent Appropriate induction rules are generated as @{thm [source] at\<^sub>f\<^sub>f.induct}, @{thm [source] ats\<^sub>f\<^sub>f.induct}, and @{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The induction rules and the underlying recursors are generated dynamically and are kept in a cache to speed up subsequent definitions. Here is a second example: \ primrec sum_btree :: "('a::{zero,plus}) btree \ 'a" and sum_btree_option :: "'a btree option \ 'a" where "sum_btree (BNode a lt rt) = a + sum_btree_option lt + sum_btree_option rt" | "sum_btree_option None = 0" | "sum_btree_option (Some t) = sum_btree t" text \ % * can pretend a nested type is mutually recursive (if purely inductive) % * avoids the higher-order map % * e.g. % * this can always be avoided; % * e.g. in our previous example, we first mapped the recursive % calls, then we used a generic at function to retrieve the result % % * there's no hard-and-fast rule of when to use one or the other, % just like there's no rule when to use fold and when to use % primrec % % * higher-order approach, considering nesting as nesting, is more % compositional -- e.g. we saw how we could reuse an existing polymorphic % at or the_default, whereas \<^const>\ats\<^sub>f\<^sub>f\ is much more specific % % * but: % * is perhaps less intuitive, because it requires higher-order thinking % * may seem inefficient, and indeed with the code generator the % mutually recursive version might be nicer % * is somewhat indirect -- must apply a map first, then compute a result % (cannot mix) % * the auxiliary functions like \<^const>\ats\<^sub>f\<^sub>f\ are sometimes useful in own right % % * impact on automation unclear % \ (*<*) end (*>*) subsection \Command Syntax \label{ssec:primrec-command-syntax}\ subsubsection \\keyw{primrec} \label{sssec:primrec}\ text \ \begin{matharray}{rcl} @{command_def "primrec"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command primrec} target? @{syntax pr_options}? fixes \ @'where' (@{syntax pr_equation} + '|') ; @{syntax_def pr_options}: '(' ((@{syntax plugins} | 'nonexhaustive' | 'transfer') + ',') ')' ; @{syntax_def pr_equation}: thmdecl? prop \ \medskip \noindent The @{command primrec} command introduces a set of mutually recursive functions over datatypes. The syntactic entity \synt{target} can be used to specify a local context, \synt{fixes} denotes a list of names with optional type signatures, \synt{thmdecl} denotes an optional name for the formula that follows, and \synt{prop} denotes a HOL proposition @{cite "isabelle-isar-ref"}. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item 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 \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 %%% TODO: elaborate on mutual and nested-to-mutual \ subsection \Generated Theorems \label{ssec:primrec-generated-theorems}\ (*<*) context early begin (*>*) text \ The @{command primrec} command generates the following properties (listed for \<^const>\tfold\): \begin{indentblock} \begin{description} \item[\f.\\hthm{simps} \[simp, code]\\rm:] ~ \\ @{thm tfold.simps(1)[no_vars]} \\ @{thm tfold.simps(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\f.\\hthm{transfer} \[transfer_rule]\\rm:] ~ \\ @{thm tfold.transfer[no_vars]} \\ This theorem is generated by the \transfer\ plugin (Section~\ref{ssec:transfer}) for functions declared with the \transfer\ option enabled. \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[\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 simultaneously. \end{description} \end{indentblock} \ (*<*) end (*>*) subsection \Recursive Default Values for Selectors \label{ssec:primrec-recursive-default-values-for-selectors}\ text \ 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 selectors are introduced. Of course, we can always define the selectors manually afterward, but we then have to state and prove all the characteristic theorems ourselves instead of letting the package do it. Fortunately, there is a workaround that relies on overloading to relieve us from the tedium of manual derivations: \begin{enumerate} \setlength{\itemsep}{0pt} \item Introduce a fully unspecified constant \un_D\<^sub>0 :: 'a\ using @{command consts}. \item Define the datatype, specifying \un_D\<^sub>0\ as the selector's default value. \item 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 \un_D\ from the characteristic equations for \un_D\<^sub>0\. \end{enumerate} \noindent The following example illustrates this procedure: \ (*<*) hide_const termi (*>*) consts termi\<^sub>0 :: 'a text \\blankline\ datatype ('a, 'b) tlist = TNil (termi: 'b) | TCons (thd: 'a) (ttl: "('a, 'b) tlist") where "ttl (TNil y) = TNil y" | "termi (TCons _ xs) = termi\<^sub>0 xs" text \\blankline\ overloading termi\<^sub>0 \ "termi\<^sub>0 :: ('a, 'b) tlist \ 'b" begin primrec termi\<^sub>0 :: "('a, 'b) tlist \ 'b" where "termi\<^sub>0 (TNil y) = y" | "termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs" end text \\blankline\ lemma termi_TCons[simp]: "termi (TCons x xs) = termi xs" by (cases xs) auto subsection \Compatibility Issues \label{ssec:primrec-compatibility-issues}\ text \ The command @{command primrec}'s behavior on new-style datatypes has been designed to be highly compatible with that for old, pre-Isabelle2015 datatypes, to ease migration. There is nonetheless at least one incompatibility that may arise when porting to the new package: \begin{itemize} \setlength{\itemsep}{0pt} \item \emph{Some theorems have different names.} For $m > 1$ mutually recursive functions, \f\<^sub>1_\_f\<^sub>m.simps\ has been broken down into separate subcollections \f\<^sub>i.simps\. \end{itemize} \ section \Defining Codatatypes \label{sec:defining-codatatypes}\ text \ Codatatypes can be specified using the @{command codatatype} command. The command is first illustrated through concrete examples featuring different flavors of corecursion. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. The \emph{Archive of Formal Proofs} also includes some useful codatatypes, notably for lazy lists @{cite "lochbihler-2010"}. \ subsection \Introductory Examples \label{ssec:codatatype-introductory-examples}\ subsubsection \Simple Corecursion \label{sssec:codatatype-simple-corecursion}\ text \ Non-corecursive codatatypes coincide with the corresponding datatypes, so they are rarely used in practice. \emph{Corecursive codatatypes} have the same syntax as recursive datatypes, except for the command name. For example, here is the definition of lazy lists: \ codatatype (lset: 'a) llist = lnull: LNil | LCons (lhd: 'a) (ltl: "'a llist") for map: lmap rel: llist_all2 pred: llist_all where "ltl LNil = LNil" text \ \noindent 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: \ codatatype (sset: 'a) stream = SCons (shd: 'a) (stl: "'a stream") for map: smap rel: stream_all2 text \ \noindent Another interesting type that can be defined as a codatatype is that of the extended natural numbers: \ codatatype enat = EZero | ESucc enat text \ \noindent This type has exactly one infinite element, \ESucc (ESucc (ESucc (\)))\, that represents $\infty$. In addition, it has finite values of the form \ESucc (\ (ESucc EZero)\)\. Here is an example with many constructors: \ codatatype 'a process = Fail | Skip (cont: "'a process") | Action (prefix: 'a) (cont: "'a process") | Choice (left: "'a process") (right: "'a process") text \ \noindent Notice that the \<^const>\cont\ selector is associated with both \<^const>\Skip\ and \<^const>\Action\. \ subsubsection \Mutual Corecursion \label{sssec:codatatype-mutual-corecursion}\ text \ \noindent The example below introduces a pair of \emph{mutually corecursive} types: \ codatatype even_enat = Even_EZero | Even_ESucc odd_enat and odd_enat = Odd_ESucc even_enat subsubsection \Nested Corecursion \label{sssec:codatatype-nested-corecursion}\ text \ \noindent The next examples feature \emph{nested corecursion}: \ 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") text \\blankline\ 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") text \\blankline\ codatatype 'a sm = SM (accept: bool) (trans: "'a \ 'a sm") subsection \Command Syntax \label{ssec:codatatype-command-syntax}\ subsubsection \\keyw{codatatype} \label{sssec:codatatype}\ text \ \begin{matharray}{rcl} @{command_def "codatatype"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command codatatype} target? @{syntax dt_options}? @{syntax dt_spec} \ \medskip \noindent Definitions of codatatypes have almost exactly the same syntax as for datatypes (Section~\ref{ssec:datatype-command-syntax}). The \discs_sels\ option is superfluous because discriminators and selectors are always generated for codatatypes. \ subsection \Generated Constants \label{ssec:codatatype-generated-constants}\ text \ 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: \medskip \begin{tabular}{@ {}ll@ {}} Corecursor: & \t.corec_t\ \end{tabular} \ subsection \Generated Theorems \label{ssec:codatatype-generated-theorems}\ text \ The characteristic theorems generated by @{command codatatype} are grouped in three broad categories: \begin{itemize} \setlength{\itemsep}{0pt} \item The \emph{free constructor theorems} (Section~\ref{sssec:free-constructor-theorems}) are properties of the constructors and destructors that can be derived for any freely generated type. \item The \emph{functorial theorems} (Section~\ref{sssec:functorial-theorems}) are properties of datatypes related to their BNF nature. \item The \emph{coinductive theorems} (Section~\ref{sssec:coinductive-theorems}) are properties of datatypes related to their coinductive nature. \end{itemize} \noindent The first two categories are exactly as for datatypes. \ subsubsection \Coinductive Theorems \label{sssec:coinductive-theorems}\ text \ The coinductive theorems are listed below for \<^typ>\'a llist\: \begin{indentblock} \begin{description} \item[\begin{tabular}{@ {}l@ {}} \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@ {}} \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@ {}} \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@ {}} \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@ {}} \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 \[consumes 1]\ is generated as well. \item[\t.\\hthm{corec}\rm:] ~ \\ @{thm llist.corec(1)[no_vars]} \\ @{thm llist.corec(2)[no_vars]} \item[\t.\\hthm{corec_code} \[code]\\rm:] ~ \\ @{thm llist.corec_code[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{corec_disc}\rm:] ~ \\ @{thm llist.corec_disc(1)[no_vars]} \\ @{thm llist.corec_disc(2)[no_vars]} \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[\t.\\hthm{corec_sel} \[simp]\\rm:] ~ \\ @{thm llist.corec_sel(1)[no_vars]} \\ @{thm llist.corec_sel(2)[no_vars]} \item[\t.\\hthm{map_o_corec}\rm:] ~ \\ @{thm llist.map_o_corec[no_vars]} \item[\t.\\hthm{corec_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm llist.corec_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} \end{indentblock} \noindent For convenience, @{command codatatype} also provides the following collection: \begin{indentblock} \begin{description} \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} \ subsection \Antiquotation \label{ssec:codatatype-antiquotation}\ subsubsection \\textit{codatatype} \label{sssec:codatatype-codatatype}\ text \ The \textit{codatatype} antiquotation, written \texttt{\char`\\\char`<\char`^}\verb|codatatype>|\guilsinglleft\textit{t}\guilsinglright{} or \texttt{@}\verb|{codatatype| \textit{t}\verb|}|, where \textit{t} is a type name, expands to \LaTeX{} code for the definition of the codatatype, with each constructor listed with its argument types. For example, if \textit{t} is @{type llist}: \begin{quote} \<^codatatype>\llist\ \end{quote} \ section \Defining Primitively Corecursive Functions \label{sec:defining-primitively-corecursive-functions}\ text \ Corecursive functions can be specified using the @{command primcorec} and \keyw{prim\-corec\-ursive} commands, which support primitive corecursion. Other approaches include the more general \keyw{partial_function} command, the \keyw{corec} and \keyw{corecursive} commands, and techniques based on domains and topologies @{cite "lochbihler-hoelzl-2014"}. In this tutorial, the focus is on @{command primcorec} and @{command primcorecursive}; \keyw{corec} and \keyw{corecursive} are described in a separate tutorial @{cite "isabelle-corec"}. More examples can be found in the directories \<^dir>\~~/src/HOL/Datatype_Examples\ and \<^dir>\~~/src/HOL/Corec_Examples\. Whereas recursive functions consume datatypes one constructor at a time, corecursive functions construct codatatypes one constructor at a time. Partly reflecting a lack of agreement among proponents of coalgebraic methods, Isabelle supports three competing syntaxes for specifying a function $f$: \begin{itemize} \setlength{\itemsep}{0pt} \abovedisplayskip=.5\abovedisplayskip \belowdisplayskip=.5\belowdisplayskip \item The \emph{destructor view} specifies $f$ by implications of the form \[\\ \ is_C\<^sub>j (f x\<^sub>1 \ x\<^sub>n)\\] and equations of the form \[\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 \[\\ \ 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 \[\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. \end{itemize} All three styles are available as input syntax. Whichever syntax is chosen, characteristic theorems for all three styles are generated. %%% TODO: partial_function? E.g. for defining tail recursive function on lazy %%% lists (cf. terminal0 in TLList.thy) \ subsection \Introductory Examples \label{ssec:primcorec-introductory-examples}\ text \ Primitive corecursion is illustrated through concrete examples based on the codatatypes defined in Section~\ref{ssec:codatatype-introductory-examples}. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. The code view is favored in the examples below. Sections \ref{ssec:primrec-constructor-view} and \ref{ssec:primrec-destructor-view} present the same examples expressed using the constructor and destructor views. \ subsubsection \Simple Corecursion \label{sssec:primcorec-simple-corecursion}\ text \ Following the code view, corecursive calls are allowed on the right-hand side as long as they occur under a constructor, which itself appears either directly to the right of the equal sign or in a conditional expression: \ primcorec (*<*)(transfer) (*>*)literate :: "('a \ 'a) \ 'a \ 'a llist" where "literate g x = LCons x (literate g (g x))" text \\blankline\ primcorec siterate :: "('a \ 'a) \ 'a \ 'a stream" where "siterate g x = SCons x (siterate g (g x))" text \ \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 \[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 from also consuming such values. The following function drops every second element in a stream: \ primcorec every_snd :: "'a stream \ 'a stream" where "every_snd s = SCons (shd s) (stl (stl s))" text \ \noindent Constructs such as \let\--\in\, \if\--\then\--\else\, and \case\--\of\ may appear around constructors that guard corecursive calls: \ primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lapp xs ys = (case xs of LNil \ ys | LCons x xs' \ LCons x (lapp xs' ys))" text \ \noindent 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 easy to generate pattern-maching equations using the @{command simps_of_case} command provided by the theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\. \ simps_of_case lapp_simps: lapp.code text \ This generates the lemma collection @{thm [source] lapp_simps}: % \begin{gather*} @{thm lapp_simps(1)[no_vars]} \\ @{thm lapp_simps(2)[no_vars]} \end{gather*} % Corecursion is useful to specify not only functions but also infinite objects: \ primcorec infty :: enat where "infty = ESucc infty" text \ \noindent The example below constructs a pseudorandom process value. It takes a stream of actions (\s\), a pseudorandom function generator (\f\), and a pseudorandom seed (\n\): \ (*<*) context early begin (*>*) primcorec random_process :: "'a stream \ (int \ int) \ int \ 'a process" where "random_process s f n = (if n mod 4 = 0 then Fail else if n mod 4 = 1 then Skip (random_process s f (f n)) else if n mod 4 = 2 then Action (shd s) (random_process (stl s) f (f n)) else Choice (random_process (every_snd s) (f \ f) (f n)) (random_process (every_snd (stl s)) (f \ f) (f (f n))))" (*<*) end (*>*) text \ \noindent The main disadvantage of the code view is that the conditions are tested sequentially. This is visible in the generated theorems. The constructor and destructor views offer nonsequential alternatives. \ subsubsection \Mutual Corecursion \label{sssec:primcorec-mutual-corecursion}\ text \ The syntax for mutually corecursive functions over mutually corecursive datatypes is unsurprising: \ primcorec even_infty :: even_enat and odd_infty :: odd_enat where "even_infty = Even_ESucc odd_infty" | "odd_infty = Odd_ESucc even_infty" subsubsection \Nested Corecursion \label{sssec:primcorec-nested-corecursion}\ 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 (\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: \ primcorec iterate\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a \ 'a tree\<^sub>i\<^sub>i" where "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i g) (g x))" text \\blankline\ primcorec iterate\<^sub>i\<^sub>s :: "('a \ 'a fset) \ 'a \ 'a tree\<^sub>i\<^sub>s" where "iterate\<^sub>i\<^sub>s g x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s g) (g x))" text \ \noindent Both examples follow the usual format for constructor arguments associated with nested recursive occurrences of the datatype. Consider \<^const>\iterate\<^sub>i\<^sub>i\. The term \<^term>\g x\ constructs an \<^typ>\'a llist\ value, which is turned into an \<^typ>\'a tree\<^sub>i\<^sub>i llist\ value using \<^const>\lmap\. This format may sometimes feel artificial. The following function constructs a tree with a single, infinite branch from a stream: \ primcorec tree\<^sub>i\<^sub>i_of_stream :: "'a stream \ 'a tree\<^sub>i\<^sub>i" where "tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (lmap tree\<^sub>i\<^sub>i_of_stream (LCons (stl s) LNil))" text \ \noindent A more natural syntax, also supported by Isabelle, is to move corecursive calls under constructors: \ primcorec (*<*)(in late) (*>*)tree\<^sub>i\<^sub>i_of_stream :: "'a stream \ 'a tree\<^sub>i\<^sub>i" where "tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (LCons (tree\<^sub>i\<^sub>i_of_stream (stl s)) LNil)" text \ The next example illustrates corecursion through functions, which is a bit special. Deterministic finite automata (DFAs) are traditionally defined as 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: \ primcorec (*<*)(in early) (*>*)sm_of_dfa :: "('q \ 'a \ 'q) \ 'q set \ 'q \ 'a sm" where "sm_of_dfa \ F q = SM (q \ F) (sm_of_dfa \ F \ \ q)" text \ \noindent 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: \ primcorec sm_of_dfa :: "('q \ 'a \ 'q) \ 'q set \ 'q \ 'a sm" where "sm_of_dfa \ F q = SM (q \ F) (\a. sm_of_dfa \ F (\ q a))" text \\blankline\ primcorec empty_sm :: "'a sm" where "empty_sm = SM False (\_. empty_sm)" text \\blankline\ primcorec not_sm :: "'a sm \ 'a sm" where "not_sm M = SM (\ accept M) (\a. not_sm (trans M a))" text \\blankline\ primcorec or_sm :: "'a sm \ 'a sm \ 'a sm" where "or_sm M N = SM (accept M \ accept N) (\a. or_sm (trans M a) (trans N a))" text \ \noindent For recursion through curried $n$-ary functions, $n$ applications of \<^term>\(\)\ are necessary. The examples below illustrate the case where $n = 2$: \ codatatype ('a, 'b) sm2 = SM2 (accept2: bool) (trans2: "'a \ 'b \ ('a, 'b) sm2") text \\blankline\ primcorec (*<*)(in early) (*>*)sm2_of_dfa :: "('q \ 'a \ 'b \ 'q) \ 'q set \ 'q \ ('a, 'b) sm2" where "sm2_of_dfa \ F q = SM2 (q \ F) ((\) ((\) (sm2_of_dfa \ F)) (\ q))" text \\blankline\ primcorec sm2_of_dfa :: "('q \ 'a \ 'b \ 'q) \ 'q set \ 'q \ ('a, 'b) sm2" where "sm2_of_dfa \ F q = SM2 (q \ F) (\a b. sm2_of_dfa \ F (\ q a b))" subsubsection \Nested-as-Mutual Corecursion \label{sssec:primcorec-nested-as-mutual-corecursion}\ text \ Just as it is possible to recurse over nested recursive datatypes as if they were mutually recursive (Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to pretend that nested codatatypes are mutually corecursive. For example: \ (*<*) context late begin (*>*) primcorec iterate\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a \ 'a tree\<^sub>i\<^sub>i" and iterates\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a llist \ 'a tree\<^sub>i\<^sub>i llist" where "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" | "iterates\<^sub>i\<^sub>i g xs = (case xs of LNil \ LNil | LCons x xs' \ LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))" text \ \noindent Coinduction rules are generated as @{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 \coinduct_strong\. These rules and the underlying corecursors are generated dynamically and are kept in a cache to speed up subsequent definitions. \ (*<*) end (*>*) subsubsection \Constructor View \label{ssec:primrec-constructor-view}\ (*<*) locale ctr_view begin (*>*) text \ The constructor view is similar to the code view, but there is one separate conditional equation per constructor rather than a single unconditional equation. Examples that rely on a single constructor, such as \<^const>\literate\ and \<^const>\siterate\, are identical in both styles. Here is an example where there is a difference: \ primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lnull xs \ lnull ys \ lapp xs ys = LNil" | "_ \ lapp xs ys = LCons (lhd (if lnull xs then ys else xs)) (if xs = LNil then ltl ys else lapp (ltl xs) ys)" text \ \noindent With the constructor view, we must distinguish between the \<^const>\LNil\ and the \<^const>\LCons\ case. The condition for \<^const>\LCons\ is left implicit, as the negation of that for \<^const>\LNil\. For this example, the constructor view is slightly more involved than the code equation. Recall the code view version presented in Section~\ref{sssec:primcorec-simple-corecursion}. % TODO: \[{thm code_view.lapp.code}\] The constructor view requires us to analyze the second argument (\<^term>\ys\). The code equation generated from the constructor view also suffers from this. % TODO: \[{thm lapp.code}\] In contrast, the next example is arguably more naturally expressed in the constructor view: \ primcorec random_process :: "'a stream \ (int \ int) \ int \ 'a process" where "n mod 4 = 0 \ random_process s f n = Fail" | "n mod 4 = 1 \ random_process s f n = Skip (random_process s f (f n))" | "n mod 4 = 2 \ random_process s f n = Action (shd s) (random_process (stl s) f (f n))" | "n mod 4 = 3 \ random_process s f n = Choice (random_process (every_snd s) f (f n)) (random_process (every_snd (stl s)) f (f n))" (*<*) end (*>*) text \ \noindent Since there is no sequentiality, we can apply the equation for \<^const>\Choice\ without having first to discharge \<^term>\n mod (4::int) \ 0\, \<^term>\n mod (4::int) \ 1\, and \<^term>\n mod (4::int) \ 2\. The price to pay for this elegance is that we must discharge exclusiveness proof 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 \sequential\ option. This pushes the problem to the users of the generated properties. %Here are more examples to conclude: \ subsubsection \Destructor View \label{ssec:primrec-destructor-view}\ (*<*) locale dtr_view begin (*>*) 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 \sequential\ option. Consider the following examples: \ primcorec literate :: "('a \ 'a) \ 'a \ 'a llist" where "\ lnull (literate _ x)" | "lhd (literate _ x) = x" | "ltl (literate g x) = literate g (g x)" text \\blankline\ primcorec siterate :: "('a \ 'a) \ 'a \ 'a stream" where "shd (siterate _ x) = x" | "stl (siterate g x) = siterate g (g x)" text \\blankline\ primcorec every_snd :: "'a stream \ 'a stream" where "shd (every_snd s) = shd s" | "stl (every_snd s) = stl (stl s)" text \ \noindent The first formula in the \<^const>\literate\ specification indicates which constructor to choose. For \<^const>\siterate\ and \<^const>\every_snd\, no such formula is necessary, since the type has only one constructor. The last two formulas are equations specifying the value of the result for the relevant selectors. Corecursive calls appear directly to the right of the equal sign. Their arguments are unrestricted. The next example shows how to specify functions that rely on more than one constructor: \ primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lnull xs \ lnull ys \ lnull (lapp xs ys)" | "lhd (lapp xs ys) = lhd (if lnull xs then ys else xs)" | "ltl (lapp xs ys) = (if xs = LNil then ltl ys else lapp (ltl xs) ys)" text \ \noindent For a codatatype with $n$ constructors, it is sufficient to specify $n - 1$ discriminator formulas. The command will then assume that the remaining constructor should be taken otherwise. This can be made explicit by adding \ (*<*) end locale dtr_view2 begin primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lnull xs \ lnull ys \ lnull (lapp xs ys)" | "lhd (lapp xs ys) = lhd (if lnull xs then ys else xs)" | "ltl (lapp xs ys) = (if xs = LNil then ltl ys else lapp (ltl xs) ys)" | (*>*) "_ \ \ lnull (lapp xs ys)" text \ \noindent to the specification. The generated selector theorems are conditional. The next example illustrates how to cope with selectors defined for several constructors: \ primcorec random_process :: "'a stream \ (int \ int) \ int \ 'a process" where "n mod 4 = 0 \ random_process s f n = Fail" | "n mod 4 = 1 \ is_Skip (random_process s f n)" | "n mod 4 = 2 \ is_Action (random_process s f n)" | "n mod 4 = 3 \ is_Choice (random_process s f n)" | "cont (random_process s f n) = random_process s f (f n)" of Skip | "prefix (random_process s f n) = shd s" | "cont (random_process s f n) = random_process (stl s) f (f n)" of Action | "left (random_process s f n) = random_process (every_snd s) f (f n)" | "right (random_process s f n) = random_process (every_snd (stl s)) f (f n)" text \ \noindent Using the \of\ keyword, different equations are specified for \<^const>\cont\ depending on which constructor is selected. Here are more examples to conclude: \ primcorec even_infty :: even_enat and odd_infty :: odd_enat where "even_infty \ Even_EZero" | "un_Even_ESucc even_infty = odd_infty" | "un_Odd_ESucc odd_infty = even_infty" text \\blankline\ primcorec iterate\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a \ 'a tree\<^sub>i\<^sub>i" where "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = x" | "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = lmap (iterate\<^sub>i\<^sub>i g) (g x)" (*<*) end (*>*) subsection \Command Syntax \label{ssec:primcorec-command-syntax}\ subsubsection \\keyw{primcorec} and \keyw{primcorecursive} \label{sssec:primcorecursive-and-primcorec}\ text \ \begin{matharray}{rcl} @{command_def "primcorec"} & : & \local_theory \ local_theory\ \\ @{command_def "primcorecursive"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ (@@{command primcorec} | @@{command primcorecursive}) target? \ @{syntax pcr_options}? fixes @'where' (@{syntax pcr_formula} + '|') ; @{syntax_def pcr_options}: '(' ((@{syntax plugins} | 'sequential' | 'exhaustive' | 'transfer') + ',') ')' ; @{syntax_def pcr_formula}: thmdecl? prop (@'of' (term * ))? \ \medskip \noindent The @{command primcorec} and @{command primcorecursive} commands introduce a set of mutually corecursive functions over codatatypes. The syntactic entity \synt{target} can be used to specify a local context, \synt{fixes} denotes a list of names with optional type signatures, \synt{thmdecl} denotes an optional name for the formula that follows, and \synt{prop} denotes a HOL proposition @{cite "isabelle-isar-ref"}. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item The \sequential\ option indicates that the conditions in specifications expressed using the constructor or destructor view are to be interpreted sequentially. \item 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 \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 \by auto?\ to discharge any emerging proof obligations. %%% TODO: elaborate on format of the propositions %%% TODO: elaborate on mutual and nested-to-mutual \ subsection \Generated Theorems \label{ssec:primcorec-generated-theorems}\ text \ The @{command primcorec} and @{command primcorecursive} commands generate the following properties (listed for \<^const>\literate\): \begin{indentblock} \begin{description} \item[\f.\\hthm{code} \[code]\\rm:] ~ \\ @{thm literate.code[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\f.\\hthm{ctr}\rm:] ~ \\ @{thm literate.ctr[no_vars]} \item[\f.\\hthm{disc} \[simp, code]\\rm:] ~ \\ @{thm literate.disc[no_vars]} \\ 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 \exhaustive\ option or whose conditions are trivially exhaustive. \item[\f.\\hthm{sel} \[simp, code]\\rm:] ~ \\ @{thm literate.disc[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \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[\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@ {}} \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@ {}} \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@ {}} \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$ mutually corecursive functions, this rule can be used to prove $m$ properties simultaneously. \item[\begin{tabular}{@ {}l@ {}} \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$ mutually corecursive functions, this rule can be used to prove $m$ properties simultaneously. \end{description} \end{indentblock} \noindent For convenience, @{command primcorec} and @{command primcorecursive} also provide the following collection: \begin{indentblock} \begin{description} \item[\f.\\hthm{simps}] = \f.disc_iff\ (or \f.disc\) \t.sel\ \end{description} \end{indentblock} \ (* subsection \Recursive Default Values for Selectors \label{ssec:primcorec-recursive-default-values-for-selectors}\ text \ partial_function to the rescue \ *) section \Registering Bounded Natural Functors \label{sec:registering-bounded-natural-functors}\ text \ The (co)datatype package can be set up to allow nested recursion through arbitrary type constructors, as long as they adhere to the BNF requirements and are registered as BNFs. It is also possible to declare a BNF abstractly without specifying its internal structure. \ subsection \Bounded Natural Functors \label{ssec:bounded-natural-functors}\ text \ Bounded natural functors (BNFs) are a semantic criterion for where (co)re\-cur\-sion may appear on the right-hand side of an equation @{cite "traytel-et-al-2012" and "blanchette-et-al-2015-wit"}. An $n$-ary BNF is a type constructor equipped with a map function (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 \llist_all :: ('a \ bool) \ 'a llist \ bool\ extends unary predicates over elements to unary predicates over lazy lists. Similarly, its relator \llist_all2 :: ('a \ 'b \ 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\. The type constructors introduced by @{command datatype} and @{command codatatype} are automatically registered as BNFs. In addition, a number of old-style datatypes and non-free types are preregistered. Given an $n$-ary BNF, the $n$ type variables associated with set functions, and on which the map function acts, are \emph{live}; any other variables are \emph{dead}. Nested (co)recursion can only take place through live variables. \ subsection \Introductory Examples \label{ssec:bnf-introductory-examples}\ text \ The example below shows how to register a type as a BNF using the @{command bnf} command. Some of the proof obligations are best viewed with the bundle "cardinal_syntax" included. The type is simply a copy of the function space \<^typ>\'d \ 'a\, where \<^typ>\'a\ is live and \<^typ>\'d\ is dead. We introduce it together with its map function, set function, predicator, and relator. \ typedef ('d, 'a) fn = "UNIV :: ('d \ 'a) set" by simp text \\blankline\ setup_lifting type_definition_fn text \\blankline\ lift_definition map_fn :: "('a \ 'b) \ ('d, 'a) fn \ ('d, 'b) fn" is "(\)" . lift_definition set_fn :: "('d, 'a) fn \ 'a set" is range . text \\blankline\ lift_definition pred_fn :: "('a \ bool) \ ('d, 'a) fn \ bool" is "pred_fun (\_. True)" . lift_definition rel_fn :: "('a \ 'b \ bool) \ ('d, 'a) fn \ ('d, 'b) fn \ bool" is "rel_fun (=)" . text \\blankline\ bnf "('d, 'a) fn" map: map_fn sets: set_fn bd: "natLeq +c |UNIV :: 'd set|" rel: rel_fn pred: pred_fn proof - show "map_fn id = id" by transfer auto next fix f :: "'a \ 'b" and g :: "'b \ 'c" show "map_fn (g \ f) = map_fn g \ map_fn f" by transfer (auto simp add: comp_def) next fix F :: "('d, 'a) fn" and f g :: "'a \ 'b" assume "\x. x \ set_fn F \ f x = g x" then show "map_fn f F = map_fn g F" by transfer auto next fix f :: "'a \ 'b" show "set_fn \ map_fn f = (`) f \ set_fn" by transfer (auto simp add: comp_def) next show "card_order (natLeq +c |UNIV :: 'd set| )" apply (rule card_order_csum) apply (rule natLeq_card_order) by (rule card_of_card_order_on) next show "cinfinite (natLeq +c |UNIV :: 'd set| )" apply (rule cinfinite_csum) apply (rule disjI1) by (rule natLeq_cinfinite) next fix F :: "('d, 'a) fn" have "|set_fn F| \o |UNIV :: 'd set|" (is "_ \o ?U") by transfer (rule card_of_image) also have "?U \o natLeq +c ?U" by (rule ordLeq_csum2) (rule card_of_Card_order) finally show "|set_fn F| \o natLeq +c |UNIV :: 'd set|" . next fix R :: "'a \ 'b \ bool" and S :: "'b \ 'c \ bool" show "rel_fn R OO rel_fn S \ rel_fn (R OO S)" by (rule, transfer) (auto simp add: rel_fun_def) next fix R :: "'a \ 'b \ bool" show "rel_fn R = (\x y. \z. set_fn z \ {(x, y). R x y} \ map_fn fst z = x \ map_fn snd z = y)" unfolding fun_eq_iff relcompp.simps conversep.simps by transfer (force simp: rel_fun_def subset_iff) next fix P :: "'a \ bool" show "pred_fn P = (\x. Ball (set_fn x) P)" unfolding fun_eq_iff by transfer simp qed text \\blankline\ print_theorems print_bnfs text \ \noindent Using \keyw{print_theorems} and @{command print_bnfs}, we can contemplate and show the world what we have achieved. This particular example does not need any nonemptiness witness, because the one generated by default is good enough, but in general this would be necessary. See \<^file>\~~/src/HOL/Basic_BNFs.thy\, \<^file>\~~/src/HOL/Library/Countable_Set_Type.thy\, \<^file>\~~/src/HOL/Library/FSet.thy\, and \<^file>\~~/src/HOL/Library/Multiset.thy\ for further examples of BNF registration, some of which feature custom witnesses. For many typedefs and quotient types, lifting the BNF structure from the raw typ to the abstract type can be done uniformly. This is the task of the @{command lift_bnf} command. Using @{command lift_bnf}, the above registration of \<^typ>\('d, 'a) fn\ as a BNF becomes much shorter: \ (*<*) context early begin (*>*) lift_bnf (*<*)(no_warn_wits) (*>*)('d, 'a) fn by force+ (*<*) end (*>*) text \ For type copies (@{command typedef}s with \<^term>\UNIV\ as the representing set), the proof obligations are so simple that they can be discharged automatically, yielding another command, @{command copy_bnf}, which does not emit any proof obligations: \ (*<*) context late begin (*>*) copy_bnf ('d, 'a) fn (*<*) end (*>*) text \ Since record schemas are type copies, @{command copy_bnf} can be used to register them as BNFs: \ record 'a point = xval :: 'a yval :: 'a text \\blankline\ copy_bnf (*<*)(no_warn_transfer) (*>*)('a, 'z) point_ext text \ In the general case, the proof obligations generated by @{command lift_bnf} are simpler than the acual BNF properties. In particular, no cardinality reasoning is required. Consider the following type of nonempty lists: \ typedef 'a nonempty_list = "{xs :: 'a list. xs \ []}" by auto text \ The @{command lift_bnf} command requires us to prove that the set of nonempty lists is closed under the map function and the zip function. The latter only occurs implicitly in the goal, in form of the variable \<^term>\zs :: ('a \ 'b) list\. \ lift_bnf (*<*)(no_warn_wits,no_warn_transfer) (*>*)'a nonempty_list proof - fix f (*<*):: "'a \ 'c"(*>*)and xs :: "'a list" assume "xs \ {xs. xs \ []}" then show "map f xs \ {xs. xs \ []}" by (cases xs(*<*) rule: list.exhaust(*>*)) auto next fix zs :: "('a \ 'b) list" assume "map fst zs \ {xs. xs \ []}" "map snd zs \ {xs. xs \ []}" then show "\zs'\{xs. xs \ []}. set zs' \ set zs \ map fst zs' = map fst zs \ map snd zs' = map snd zs" by (cases zs (*<*)rule: list.exhaust(*>*)) (auto intro!: exI[of _ zs]) qed text \The @{command lift_bnf} command also supports quotient types. Here is an example that defines the option type as a quotient of the sum type. The proof obligations generated by @{command lift_bnf} for quotients are different from the ones for typedefs. You can find additional examples of usages of @{command lift_bnf} for both quotients and subtypes in the session \emph{HOL-Datatype_Examples}.\ inductive ignore_Inl :: "'a + 'a \ 'a + 'a \ bool" where "ignore_Inl (Inl x) (Inl y)" | "ignore_Inl (Inr x) (Inr x)" (*<*) inductive_simps ignore_Inl_simps[simp]: "ignore_Inl (Inl x) y" "ignore_Inl (Inr x') y" "ignore_Inl y (Inl x)" "ignore_Inl y (Inr x')" (*>*) lemma ignore_Inl_equivp: "ignore_Inl x x" "ignore_Inl x y \ ignore_Inl y x" "ignore_Inl x y \ ignore_Inl y z \ ignore_Inl x z" by (cases x; cases y; cases z; auto)+ quotient_type 'a myoption = "'a + 'a" / ignore_Inl unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def by (blast intro: ignore_Inl_equivp) lift_bnf 'a myoption (*<*)[wits: "Inl undefined :: 'a + 'a"](*>*) proof - fix P :: "'a \ 'b \ bool" and Q :: "'b \ 'c \ bool" assume "P OO Q \ bot" then show "rel_sum P P OO ignore_Inl OO rel_sum Q Q \ ignore_Inl OO rel_sum (P OO Q) (P OO Q) OO ignore_Inl" by (fastforce(*<*) elim!: ignore_Inl.cases simp add: relcompp_apply fun_eq_iff rel_sum.simps(*>*)) next fix S :: "'a set set" let ?eq = "{(x, x'). ignore_Inl x x'}" let ?in = "\A. {x. Basic_BNFs.setl x \ Basic_BNFs.setr x \ A}" assume "S \ {}" "\ S \ {}" show "(\A\S. ?eq `` ?in A) \ ?eq `` ?in (\ S)" proof (intro subsetI) fix x assume "x \ (\A\S. ?eq `` ?in A)" with \\ S \ {}\ show "x \ ?eq `` ?in (\ S)" by (cases x) (fastforce(*<*) dest!: spec[where x="Inl _"](*>*))+ qed (*<*)next fix a :: 'a assume "a \ (\mx\{mx. ignore_Inl (map_sum Inr Inr (Inl undefined)) mx}. \ (Basic_BNFs.setr ` (Basic_BNFs.setl mx \ Basic_BNFs.setr mx)))" then show False by (auto simp: setr.simps)(*>*) qed 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 \('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 \'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 \('a, 'b, 'c) F\. The BNF properties are postulated as axioms. \ bnf_axiomatization (setA: 'a, setB: 'b, setC: 'c) F [wits: "'a \ ('a, 'b, 'c) F"] text \\blankline\ print_theorems print_bnfs subsection \Command Syntax \label{ssec:bnf-command-syntax}\ subsubsection \\keyw{bnf} \label{sssec:bnf}\ text \ \begin{matharray}{rcl} @{command_def "bnf"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command bnf} target? (name ':')? type \ 'map:' term ('sets:' (term +))? 'bd:' term \ ('wits:' (term +))? ('rel:' term)? \ ('pred:' term)? @{syntax plugins}? \ \medskip \noindent The @{command bnf} command registers an existing type as a bounded natural functor (BNF). The type must be equipped with an appropriate map function (functorial action). In addition, custom set functions, predicators, relators, and nonemptiness witnesses can be specified; otherwise, default versions are used. The syntactic entity \synt{target} can be used to specify a local context, \synt{type} denotes a HOL type, and \synt{term} denotes a HOL term @{cite "isabelle-isar-ref"}. The @{syntax plugins} option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. %%% TODO: elaborate on proof obligations \ subsubsection \\keyw{lift_bnf} \label{sssec:lift-bnf}\ text \ \begin{matharray}{rcl} @{command_def "lift_bnf"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command lift_bnf} target? lb_options? \ @{syntax tyargs} name wit_terms? \ - ('via' thm)? @{syntax map_rel_pred}? + ('via' thm thm?)? @{syntax map_rel_pred}? ; @{syntax_def lb_options}: '(' ((@{syntax plugins} | 'no_warn_wits' | 'no_warn_transfer') + ',') ')' ; @{syntax_def wit_terms}: '[' 'wits' ':' terms ']' \ \medskip \noindent The @{command lift_bnf} command registers as a BNF an existing type (the \emph{abstract type}) that was defined as a subtype of a BNF (the \emph{raw type}) using the @{command typedef} command or as a quotient type of a BNF (also, the \emph{raw type}) using the @{command quotient_type}. To achieve this, it lifts the BNF structure on the raw type to the abstract type following a \<^term>\type_definition\ or a \<^term>\Quotient\ theorem. The theorem is usually inferred from the type, but can -also be explicitly supplied by means of the optional \via\ clause. In +also be explicitly supplied by means of the optional \via\ clause. In case of quotients, it +is sometimes also necessary to supply a second theorem of the form \<^term>\reflp eq\, +that expresses the reflexivity (and thus totality) of the equivalence relation. In addition, custom names for the set functions, the map function, the predicator, and the relator, as well as nonemptiness witnesses can be specified. Nonemptiness witnesses are not lifted from the raw type's BNF, as this would be 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 \no_warn_wits\ option. \ subsubsection \\keyw{copy_bnf} \label{sssec:copy-bnf}\ text \ \begin{matharray}{rcl} @{command_def "copy_bnf"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command copy_bnf} target? cb_options? \ @{syntax tyargs} name ('via' thm)? @{syntax map_rel_pred}? ; @{syntax_def cb_options}: '(' ((@{syntax plugins} | 'no_warn_transfer') + ',') ')' \ \medskip \noindent The @{command copy_bnf} command performs the same lifting as @{command lift_bnf} for type copies (@{command typedef}s with \<^term>\UNIV\ as the representing set), without requiring the user to discharge any proof obligations or provide nonemptiness witnesses. \ subsubsection \\keyw{bnf_axiomatization} \label{sssec:bnf-axiomatization}\ text \ \begin{matharray}{rcl} @{command_def "bnf_axiomatization"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command bnf_axiomatization} target? ('(' @{syntax plugins} ')')? \ @{syntax tyargs}? name @{syntax wit_types}? \ mixfix? @{syntax map_rel_pred}? ; @{syntax_def wit_types}: '[' 'wits' ':' types ']' \ \medskip \noindent The @{command bnf_axiomatization} command declares a new type and associated constants (map, set, predicator, relator, and cardinal bound) and asserts the BNF properties for these constants as axioms. The syntactic entity \synt{target} can be used to specify a local context, \synt{name} denotes an identifier, \synt{typefree} denotes fixed type variable (\<^typ>\'a\, \<^typ>\'b\, \ldots), \synt{mixfix} denotes the usual parenthesized mixfix notation, and \synt{types} denotes a space-separated list of types @{cite "isabelle-isar-ref"}. The @{syntax plugins} option indicates which plugins should be 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 \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 @{command codatatype} definition. The command is useful to reason abstractly about BNFs. The axioms are safe because there exist BNFs of arbitrary large arities. Applications must import the \<^file>\~~/src/HOL/Library/BNF_Axiomatization.thy\ theory to use this functionality. \ subsubsection \\keyw{print_bnfs} \label{sssec:print-bnfs}\ text \ \begin{matharray}{rcl} @{command_def "print_bnfs"} & : & \local_theory \\ \end{matharray} \<^rail>\ @@{command print_bnfs} \ \ section \Deriving Destructors and Theorems for Free Constructors \label{sec:deriving-destructors-and-theorems-for-free-constructors}\ text \ The derivation of convenience theorems for types equipped with free constructors, as performed internally by @{command datatype} and @{command codatatype}, is available as a stand-alone command called @{command free_constructors}. % * need for this is rare but may arise if you want e.g. to add destructors to % a type not introduced by ... % % * @{command free_constructors} % * \plugins\, \discs_sels\ % * hack to have both co and nonco view via locale (cf. ext nats) \ (* subsection \Introductory Example \label{ssec:ctors-introductory-example}\ *) subsection \Command Syntax \label{ssec:ctors-command-syntax}\ subsubsection \\keyw{free_constructors} \label{sssec:free-constructors}\ text \ \begin{matharray}{rcl} @{command_def "free_constructors"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command free_constructors} target? @{syntax dt_options} \ name 'for' (@{syntax fc_ctor} + '|') \ (@'where' (prop + '|'))? ; @{syntax_def fc_ctor}: (name ':')? term (name * ) \ \medskip \noindent The @{command free_constructors} command generates destructor constants for freely constructed types as well as properties about constructors and destructors. It also registers the constants and theorems in a data structure that is queried by various tools (e.g., \keyw{function}). The syntactic entity \synt{target} can be used to specify a local context, \synt{name} denotes an identifier, \synt{prop} denotes a HOL proposition, and \synt{term} denotes a HOL term @{cite "isabelle-isar-ref"}. The syntax resembles that of @{command datatype} and @{command codatatype} definitions (Sections \ref{ssec:datatype-command-syntax}~and~\ref{ssec:codatatype-command-syntax}). A constructor is specified by an optional name for the discriminator, the 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 \[fundef_cong]\ attribute is not set on the generated \case_cong\ theorem. It can be added manually using \keyw{declare}. \ subsubsection \\keyw{simps_of_case} \label{sssec:simps-of-case}\ text \ \begin{matharray}{rcl} @{command_def "simps_of_case"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command simps_of_case} target? (name ':')? \ (thm + ) (@'splits' ':' (thm + ))? \ \medskip \noindent The @{command simps_of_case} command provided by theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\ converts a single equation with a complex case expression on the right-hand side into a set of pattern-matching equations. For example, \ (*<*) context late begin (*>*) simps_of_case lapp_simps: lapp.code text \ \noindent translates @{thm lapp.code[no_vars]} into % \begin{gather*} @{thm lapp_simps(1)[no_vars]} \\ @{thm lapp_simps(2)[no_vars]} \end{gather*} \ subsubsection \\keyw{case_of_simps} \label{sssec:case-of-simps}\ text \ \begin{matharray}{rcl} @{command_def "case_of_simps"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command case_of_simps} target? (name ':')? \ (thm + ) \ \medskip \noindent The \keyw{case_of_simps} command provided by theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\ converts a set of pattern-matching equations into single equation with a complex case expression on the right-hand side (cf.\ @{command simps_of_case}). For example, \ case_of_simps lapp_case: lapp_simps text \ \noindent translates % \begin{gather*} @{thm lapp_simps(1)[no_vars]} \\ @{thm lapp_simps(2)[no_vars]} \end{gather*} % into @{thm lapp_case[no_vars]}. \ (*<*) end (*>*) (* section \Using the Standard ML Interface \label{sec:using-the-standard-ml-interface}\ text \ The package's programmatic interface. \ *) section \Selecting Plugins \label{sec:selecting-plugins}\ text \ Plugins extend the (co)datatype package to interoperate with other Isabelle packages and tools, such as the code generator, Transfer, Lifting, and Quickcheck. They can be enabled or disabled individually using the @{syntax plugins} option to the commands @{command datatype}, @{command primrec}, @{command codatatype}, @{command primcorec}, @{command primcorecursive}, @{command bnf}, @{command bnf_axiomatization}, and @{command free_constructors}. For example: \ datatype (plugins del: code "quickcheck") color = Red | Black text \ Beyond the standard plugins, the \emph{Archive of Formal Proofs} includes a \keyw{derive} command that derives class instances of datatypes @{cite "sternagel-thiemann-2015"}. \ subsection \Code Generator \label{ssec:code-generator}\ text \ The \hthm{code} plugin registers freely generated types, including (co)datatypes, and (co)recursive functions for code generation. No distinction is made between datatypes and codatatypes. This means that for target languages with a strict evaluation strategy (e.g., Standard ML), programs that attempt to produce infinite codatatype values will not terminate. For types, the plugin derives the following properties: \begin{indentblock} \begin{description} \item[\t.\\hthm{eq.refl} \[code nbe]\\rm:] ~ \\ @{thm list.eq.refl[no_vars]} \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]} \\ @{thm list.eq.simps(4)[no_vars]} \\ @{thm list.eq.simps(5)[no_vars]} \\ @{thm list.eq.simps(6)[no_vars]} \end{description} \end{indentblock} 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}, and~\ref{ssec:primcorec-generated-theorems}. \ subsection \Size \label{ssec:size}\ text \ 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. The plugin derives the following properties: \begin{indentblock} \begin{description} \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[\t.\\hthm{size_gen}\rm:] ~ \\ @{thm list.size_gen(1)[no_vars]} \\ @{thm list.size_gen(2)[no_vars]} \item[\t.\\hthm{size_gen_o_map}\rm:] ~ \\ @{thm list.size_gen_o_map[no_vars]} \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\. \end{description} \end{indentblock} 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 \('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. \ subsection \Transfer \label{ssec:transfer}\ text \ For each (co)datatype with live type arguments and each manually registered BNF, 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 arguments}, the plugin derives the following properties: \begin{indentblock} \begin{description} \item[\t.\\hthm{Domainp_rel} \[relator_domain]\\rm:] ~ \\ @{thm list.Domainp_rel[no_vars]} \item[\t.\\hthm{left_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.left_total_rel[no_vars]} \item[\t.\\hthm{left_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.left_unique_rel[no_vars]} \item[\t.\\hthm{right_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.right_total_rel[no_vars]} \item[\t.\\hthm{right_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.right_unique_rel[no_vars]} \item[\t.\\hthm{bi_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.bi_total_rel[no_vars]} \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 \[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 \[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 \f.transfer\ property, conditioned by the \transfer\ option, and sets the \[transfer_rule]\ attribute on these. \ subsection \Lifting \label{ssec:lifting}\ text \ For each (co)datatype and each manually registered BNF with at least one live type argument \emph{and no dead type arguments}, the \hthm{lifting} plugin generates properties and attributes that guide the Lifting tool. The plugin derives the following property: \begin{indentblock} \begin{description} \item[\t.\\hthm{Quotient} \[quot_map]\\rm:] ~ \\ @{thm list.Quotient[no_vars]} \end{description} \end{indentblock} 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\. \ subsection \Quickcheck \label{ssec:quickcheck}\ text \ The integration of datatypes with Quickcheck is accomplished by the \hthm{quick\-check} plugin. It combines a number of subplugins that instantiate specific type classes. The subplugins can be enabled or disabled individually. They are listed below: \begin{indentblock} \hthm{quickcheck_random} \\ \hthm{quickcheck_exhaustive} \\ \hthm{quickcheck_bounded_forall} \\ \hthm{quickcheck_full_exhaustive} \\ \hthm{quickcheck_narrowing} \end{indentblock} \ subsection \Program Extraction \label{ssec:program-extraction}\ text \ The \hthm{extraction} plugin provides realizers for induction and case analysis, to enable program extraction from proofs involving datatypes. This functionality is only available with full proof objects, i.e., with the \emph{HOL-Proofs} session. \ section \Known Bugs and Limitations \label{sec:known-bugs-and-limitations}\ text \ This section lists the known bugs and limitations of the (co)datatype package at the time of this writing. \begin{enumerate} \setlength{\itemsep}{0pt} \item \emph{Defining mutually (co)recursive (co)datatypes can be slow.} Fortunately, it is always possible to recast mutual specifications to nested ones, which are processed more efficiently. \item \emph{Locally fixed types and terms cannot be used in type specifications.} The limitation on types can be circumvented by adding type arguments to the local (co)datatypes to abstract over the locally fixed types. \item \emph{The \emph{\keyw{primcorec}} command does not allow user-specified names and attributes next to the entered formulas.} The less convenient syntax, using the \keyw{lemmas} command, is available as an alternative. \item \emph{The \emph{\keyw{primcorec}} command does not allow corecursion under \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 \0\, as a constructor.} \item \emph{There is no way to register the same type as both a datatype and a codatatype.} This affects types such as the extended natural numbers, for which both views would make sense (for a different set of constructors). \item \emph{The names of variables are often suboptimal in the properties generated by the package.} \item \emph{The compatibility layer sometimes produces induction principles with a slightly different ordering of the premises than the old package.} \end{enumerate} \ text \ \section*{Acknowledgment} Tobias Nipkow and Makarius Wenzel encouraged us to implement the new (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 \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 case_of_simps} commands. Florian Haftmann, Christian Urban, and Makarius Wenzel provided general advice on Isabelle and package writing. Stefan Milius and Lutz Schr\"oder found an elegant proof that eliminated one of the BNF proof obligations. Mamoun Filali-Amine, Gerwin Klein, Andreas Lochbihler, Tobias Nipkow, and Christian Sternagel suggested many textual improvements to this tutorial. \ end diff --git a/src/HOL/Datatype_Examples/Dlist.thy b/src/HOL/Datatype_Examples/Dlist.thy deleted file mode 100644 --- a/src/HOL/Datatype_Examples/Dlist.thy +++ /dev/null @@ -1,81 +0,0 @@ -theory Dlist imports - "HOL-Library.Confluent_Quotient" -begin - -context begin - -qualified definition dlist_eq where "dlist_eq = BNF_Def.vimage2p remdups remdups (=)" - -lemma equivp_dlist_eq: "equivp dlist_eq" - unfolding dlist_eq_def by(rule equivp_vimage2p)(rule identity_equivp) - -quotient_type 'a dlist = "'a list" / dlist_eq - by (rule equivp_dlist_eq) - -qualified inductive double :: "'a list \ 'a list \ bool" where - "double (xs @ ys) (xs @ x # ys)" if "x \ set ys" - -qualified lemma strong_confluentp_double: "strong_confluentp double" -proof - fix xs ys zs :: "'a list" - assume ys: "double xs ys" and zs: "double xs zs" - consider (left) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ y # bs @ cs" "zs = as @ bs @ z # cs" "y \ set (bs @ cs)" "z \ set cs" - | (right) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ bs @ y # cs" "zs = as @ z # bs @ cs" "y \ set cs" "z \ set (bs @ cs)" - proof - - show thesis using ys zs - by(clarsimp simp add: double.simps append_eq_append_conv2)(auto intro: that) - qed - then show "\us. double\<^sup>*\<^sup>* ys us \ double\<^sup>=\<^sup>= zs us" - proof cases - case left - let ?us = "as @ y # bs @ z # cs" - have "double ys ?us" "double zs ?us" using left - by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+ - then show ?thesis by blast - next - case right - let ?us = "as @ z # bs @ y # cs" - have "double ys ?us" "double zs ?us" using right - by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+ - then show ?thesis by blast - qed -qed - -qualified lemma double_Cons1 [simp]: "double xs (x # xs)" if "x \ set xs" - using double.intros[of x xs "[]"] that by simp - -qualified lemma double_Cons_same [simp]: "double xs ys \ double (x # xs) (x # ys)" - by(auto simp add: double.simps Cons_eq_append_conv) - -qualified lemma doubles_Cons_same: "double\<^sup>*\<^sup>* xs ys \ double\<^sup>*\<^sup>* (x # xs) (x # ys)" - by(induction rule: rtranclp_induct)(auto intro: rtranclp.rtrancl_into_rtrancl) - -qualified lemma remdups_into_doubles: "double\<^sup>*\<^sup>* (remdups xs) xs" - by(induction xs)(auto intro: doubles_Cons_same rtranclp.rtrancl_into_rtrancl) - -qualified lemma dlist_eq_into_doubles: "dlist_eq \ equivclp double" - by(auto 4 4 simp add: dlist_eq_def vimage2p_def - intro: equivclp_trans converse_rtranclp_into_equivclp rtranclp_into_equivclp remdups_into_doubles) - -qualified lemma factor_double_map: "double (map f xs) ys \ \zs. dlist_eq xs zs \ ys = map f zs" - by(auto simp add: double.simps dlist_eq_def vimage2p_def map_eq_append_conv) - (metis list.simps(9) map_append remdups.simps(2) remdups_append2) - -qualified lemma dlist_eq_set_eq: "dlist_eq xs ys \ set xs = set ys" - by(simp add: dlist_eq_def vimage2p_def)(metis set_remdups) - -qualified lemma dlist_eq_map_respect: "dlist_eq xs ys \ dlist_eq (map f xs) (map f ys)" - by(clarsimp simp add: dlist_eq_def vimage2p_def)(metis remdups_map_remdups) - -qualified lemma confluent_quotient_dlist: - "confluent_quotient double dlist_eq dlist_eq dlist_eq dlist_eq dlist_eq (map fst) (map snd) (map fst) (map snd) list_all2 list_all2 list_all2 set set" - by(unfold_locales)(auto intro: strong_confluentp_imp_confluentp strong_confluentp_double dest: factor_double_map dlist_eq_into_doubles[THEN predicate2D] dlist_eq_set_eq simp add: list.in_rel list.rel_compp dlist_eq_map_respect equivp_dlist_eq equivp_imp_transp) - -lift_bnf 'a dlist - subgoal for A B by(rule confluent_quotient.subdistributivity[OF confluent_quotient_dlist]) - subgoal by(force dest: dlist_eq_set_eq intro: equivp_reflp[OF equivp_dlist_eq]) - done - -end - -end diff --git a/src/HOL/Library/Dlist.thy b/src/HOL/Library/Dlist.thy --- a/src/HOL/Library/Dlist.thy +++ b/src/HOL/Library/Dlist.thy @@ -1,344 +1,323 @@ -(* Author: Florian Haftmann, TU Muenchen +(* Author: Florian Haftmann, TU Muenchen Author: Andreas Lochbihler, ETH Zurich *) section \Lists with elements distinct as canonical example for datatype invariants\ theory Dlist -imports Main +imports Confluent_Quotient begin subsection \The type of distinct lists\ typedef 'a dlist = "{xs::'a list. distinct xs}" morphisms list_of_dlist Abs_dlist proof show "[] \ {xs. distinct xs}" by simp qed +context begin + +qualified definition dlist_eq where "dlist_eq = BNF_Def.vimage2p remdups remdups (=)" + +qualified lemma equivp_dlist_eq: "equivp dlist_eq" + unfolding dlist_eq_def by(rule equivp_vimage2p)(rule identity_equivp) + +qualified definition abs_dlist :: "'a list \ 'a dlist" where "abs_dlist = Abs_dlist o remdups" + +definition qcr_dlist :: "'a list \ 'a dlist \ bool" where "qcr_dlist x y \ y = abs_dlist x" + +qualified lemma Quotient_dlist_remdups: "Quotient dlist_eq abs_dlist list_of_dlist qcr_dlist" + unfolding Quotient_def dlist_eq_def qcr_dlist_def vimage2p_def abs_dlist_def + by (auto simp add: fun_eq_iff Abs_dlist_inject + list_of_dlist[simplified] list_of_dlist_inverse distinct_remdups_id) + +end + +locale Quotient_dlist begin +setup_lifting Dlist.Quotient_dlist_remdups Dlist.equivp_dlist_eq[THEN equivp_reflp2] +end + setup_lifting type_definition_dlist lemma dlist_eq_iff: "dxs = dys \ list_of_dlist dxs = list_of_dlist dys" by (simp add: list_of_dlist_inject) lemma dlist_eqI: "list_of_dlist dxs = list_of_dlist dys \ dxs = dys" by (simp add: dlist_eq_iff) text \Formal, totalized constructor for \<^typ>\'a dlist\:\ definition Dlist :: "'a list \ 'a dlist" where "Dlist xs = Abs_dlist (remdups xs)" lemma distinct_list_of_dlist [simp, intro]: "distinct (list_of_dlist dxs)" using list_of_dlist [of dxs] by simp lemma list_of_dlist_Dlist [simp]: "list_of_dlist (Dlist xs) = remdups xs" by (simp add: Dlist_def Abs_dlist_inverse) lemma remdups_list_of_dlist [simp]: "remdups (list_of_dlist dxs) = list_of_dlist dxs" by simp lemma Dlist_list_of_dlist [simp, code abstype]: "Dlist (list_of_dlist dxs) = dxs" by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id) text \Fundamental operations:\ context begin qualified definition empty :: "'a dlist" where "empty = Dlist []" qualified definition insert :: "'a \ 'a dlist \ 'a dlist" where "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))" qualified definition remove :: "'a \ 'a dlist \ 'a dlist" where "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))" qualified definition map :: "('a \ 'b) \ 'a dlist \ 'b dlist" where "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))" qualified definition filter :: "('a \ bool) \ 'a dlist \ 'a dlist" where "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))" qualified definition rotate :: "nat \ 'a dlist \ 'a dlist" where "rotate n dxs = Dlist (List.rotate n (list_of_dlist dxs))" end text \Derived operations:\ context begin qualified definition null :: "'a dlist \ bool" where "null dxs = List.null (list_of_dlist dxs)" qualified definition member :: "'a dlist \ 'a \ bool" where "member dxs = List.member (list_of_dlist dxs)" qualified definition length :: "'a dlist \ nat" where "length dxs = List.length (list_of_dlist dxs)" qualified definition fold :: "('a \ 'b \ 'b) \ 'a dlist \ 'b \ 'b" where "fold f dxs = List.fold f (list_of_dlist dxs)" qualified definition foldr :: "('a \ 'b \ 'b) \ 'a dlist \ 'b \ 'b" where "foldr f dxs = List.foldr f (list_of_dlist dxs)" end subsection \Executable version obeying invariant\ lemma list_of_dlist_empty [simp, code abstract]: "list_of_dlist Dlist.empty = []" by (simp add: Dlist.empty_def) lemma list_of_dlist_insert [simp, code abstract]: "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)" by (simp add: Dlist.insert_def) lemma list_of_dlist_remove [simp, code abstract]: "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)" by (simp add: Dlist.remove_def) lemma list_of_dlist_map [simp, code abstract]: "list_of_dlist (Dlist.map f dxs) = remdups (List.map f (list_of_dlist dxs))" by (simp add: Dlist.map_def) lemma list_of_dlist_filter [simp, code abstract]: "list_of_dlist (Dlist.filter P dxs) = List.filter P (list_of_dlist dxs)" by (simp add: Dlist.filter_def) lemma list_of_dlist_rotate [simp, code abstract]: "list_of_dlist (Dlist.rotate n dxs) = List.rotate n (list_of_dlist dxs)" by (simp add: Dlist.rotate_def) text \Explicit executable conversion\ definition dlist_of_list [simp]: "dlist_of_list = Dlist" lemma [code abstract]: "list_of_dlist (dlist_of_list xs) = remdups xs" by simp text \Equality\ instantiation dlist :: (equal) equal begin definition "HOL.equal dxs dys \ HOL.equal (list_of_dlist dxs) (list_of_dlist dys)" instance by standard (simp add: equal_dlist_def equal list_of_dlist_inject) end declare equal_dlist_def [code] lemma [code nbe]: "HOL.equal (dxs :: 'a::equal dlist) dxs \ True" by (fact equal_refl) subsection \Induction principle and case distinction\ lemma dlist_induct [case_names empty insert, induct type: dlist]: assumes empty: "P Dlist.empty" assumes insrt: "\x dxs. \ Dlist.member dxs x \ P dxs \ P (Dlist.insert x dxs)" shows "P dxs" proof (cases dxs) case (Abs_dlist xs) then have "distinct xs" and dxs: "dxs = Dlist xs" by (simp_all add: Dlist_def distinct_remdups_id) from \distinct xs\ have "P (Dlist xs)" proof (induct xs) case Nil from empty show ?case by (simp add: Dlist.empty_def) next case (Cons x xs) then have "\ Dlist.member (Dlist xs) x" and "P (Dlist xs)" by (simp_all add: Dlist.member_def List.member_def) with insrt have "P (Dlist.insert x (Dlist xs))" . with Cons show ?case by (simp add: Dlist.insert_def distinct_remdups_id) qed with dxs show "P dxs" by simp qed lemma dlist_case [cases type: dlist]: obtains (empty) "dxs = Dlist.empty" | (insert) x dys where "\ Dlist.member dys x" and "dxs = Dlist.insert x dys" proof (cases dxs) case (Abs_dlist xs) then have dxs: "dxs = Dlist xs" and distinct: "distinct xs" by (simp_all add: Dlist_def distinct_remdups_id) show thesis proof (cases xs) case Nil with dxs - have "dxs = Dlist.empty" by (simp add: Dlist.empty_def) + have "dxs = Dlist.empty" by (simp add: Dlist.empty_def) with empty show ?thesis . next case (Cons x xs) with dxs distinct have "\ Dlist.member (Dlist xs) x" and "dxs = Dlist.insert x (Dlist xs)" by (simp_all add: Dlist.member_def List.member_def Dlist.insert_def distinct_remdups_id) with insert show ?thesis . qed qed subsection \Functorial structure\ functor map: map by (simp_all add: remdups_map_remdups fun_eq_iff dlist_eq_iff) subsection \Quickcheck generators\ quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert subsection \BNF instance\ context begin -qualified fun wpull :: "('a \ 'b) list \ ('b \ 'c) list \ ('a \ 'c) list" -where - "wpull [] ys = []" -| "wpull xs [] = []" -| "wpull ((a, b) # xs) ((b', c) # ys) = - (if b \ snd ` set xs then - (a, the (map_of (rev ((b', c) # ys)) b)) # wpull xs ((b', c) # ys) - else if b' \ fst ` set ys then - (the (map_of (map prod.swap (rev ((a, b) # xs))) b'), c) # wpull ((a, b) # xs) ys - else (a, c) # wpull xs ys)" - -qualified lemma wpull_eq_Nil_iff [simp]: "wpull xs ys = [] \ xs = [] \ ys = []" -by(cases "(xs, ys)" rule: wpull.cases) simp_all +qualified inductive double :: "'a list \ 'a list \ bool" where + "double (xs @ ys) (xs @ x # ys)" if "x \ set ys" -qualified lemma wpull_induct - [consumes 1, - case_names Nil left[xs eq in_set IH] right[xs ys eq in_set IH] step[xs ys eq IH] ]: - assumes eq: "remdups (map snd xs) = remdups (map fst ys)" - and Nil: "P [] []" - and left: "\a b xs b' c ys. - \ b \ snd ` set xs; remdups (map snd xs) = remdups (map fst ((b', c) # ys)); - (b, the (map_of (rev ((b', c) # ys)) b)) \ set ((b', c) # ys); P xs ((b', c) # ys) \ - \ P ((a, b) # xs) ((b', c) # ys)" - and right: "\a b xs b' c ys. - \ b \ snd ` set xs; b' \ fst ` set ys; - remdups (map snd ((a, b) # xs)) = remdups (map fst ys); - (the (map_of (map prod.swap (rev ((a, b) #xs))) b'), b') \ set ((a, b) # xs); - P ((a, b) # xs) ys \ - \ P ((a, b) # xs) ((b', c) # ys)" - and step: "\a b xs c ys. - \ b \ snd ` set xs; b \ fst ` set ys; remdups (map snd xs) = remdups (map fst ys); - P xs ys \ - \ P ((a, b) # xs) ((b, c) # ys)" - shows "P xs ys" -using eq -proof(induction xs ys rule: wpull.induct) - case 1 thus ?case by(simp add: Nil) -next - case 2 thus ?case by(simp split: if_split_asm) -next - case Cons: (3 a b xs b' c ys) - let ?xs = "(a, b) # xs" and ?ys = "(b', c) # ys" - consider (xs) "b \ snd ` set xs" | (ys) "b \ snd ` set xs" "b' \ fst ` set ys" - | (step) "b \ snd ` set xs" "b' \ fst ` set ys" by auto - thus ?case +qualified lemma strong_confluentp_double: "strong_confluentp double" +proof + fix xs ys zs :: "'a list" + assume ys: "double xs ys" and zs: "double xs zs" + consider (left) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ y # bs @ cs" "zs = as @ bs @ z # cs" "y \ set (bs @ cs)" "z \ set cs" + | (right) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ bs @ y # cs" "zs = as @ z # bs @ cs" "y \ set cs" "z \ set (bs @ cs)" + proof - + show thesis using ys zs + by(clarsimp simp add: double.simps append_eq_append_conv2)(auto intro: that) + qed + then show "\us. double\<^sup>*\<^sup>* ys us \ double\<^sup>=\<^sup>= zs us" proof cases - case xs - with Cons.prems have eq: "remdups (map snd xs) = remdups (map fst ?ys)" by auto - from xs eq have "b \ fst ` set ?ys" by (metis list.set_map set_remdups) - hence "map_of (rev ?ys) b \ None" unfolding map_of_eq_None_iff by auto - then obtain c' where "map_of (rev ?ys) b = Some c'" by blast - then have "(b, the (map_of (rev ?ys) b)) \ set ?ys" by(auto dest: map_of_SomeD split: if_split_asm) - from xs eq this Cons.IH(1)[OF xs eq] show ?thesis by(rule left) + case left + let ?us = "as @ y # bs @ z # cs" + have "double ys ?us" "double zs ?us" using left + by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+ + then show ?thesis by blast next - case ys - from ys Cons.prems have eq: "remdups (map snd ?xs) = remdups (map fst ys)" by auto - from ys eq have "b' \ snd ` set ?xs" by (metis list.set_map set_remdups) - hence "map_of (map prod.swap (rev ?xs)) b' \ None" - unfolding map_of_eq_None_iff by(auto simp add: image_image) - then obtain a' where "map_of (map prod.swap (rev ?xs)) b' = Some a'" by blast - then have "(the (map_of (map prod.swap (rev ?xs)) b'), b') \ set ?xs" - by(auto dest: map_of_SomeD split: if_split_asm) - from ys eq this Cons.IH(2)[OF ys eq] show ?thesis by(rule right) - next - case *: step - hence "remdups (map snd xs) = remdups (map fst ys)" "b = b'" using Cons.prems by auto - from * this(1) Cons.IH(3)[OF * this(1)] show ?thesis unfolding \b = b'\ by(rule step) + case right + let ?us = "as @ z # bs @ y # cs" + have "double ys ?us" "double zs ?us" using right + by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+ + then show ?thesis by blast qed qed -qualified lemma set_wpull_subset: - assumes "remdups (map snd xs) = remdups (map fst ys)" - shows "set (wpull xs ys) \ set xs O set ys" -using assms by(induction xs ys rule: wpull_induct) auto - -qualified lemma set_fst_wpull: - assumes "remdups (map snd xs) = remdups (map fst ys)" - shows "fst ` set (wpull xs ys) = fst ` set xs" -using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI) +qualified lemma double_Cons1 [simp]: "double xs (x # xs)" if "x \ set xs" + using double.intros[of x xs "[]"] that by simp -qualified lemma set_snd_wpull: - assumes "remdups (map snd xs) = remdups (map fst ys)" - shows "snd ` set (wpull xs ys) = snd ` set ys" -using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI) - -qualified lemma wpull: - assumes "distinct xs" - and "distinct ys" - and "set xs \ {(x, y). R x y}" - and "set ys \ {(x, y). S x y}" - and eq: "remdups (map snd xs) = remdups (map fst ys)" - shows "\zs. distinct zs \ set zs \ {(x, y). (R OO S) x y} \ - remdups (map fst zs) = remdups (map fst xs) \ remdups (map snd zs) = remdups (map snd ys)" -proof(intro exI conjI) - let ?zs = "remdups (wpull xs ys)" - show "distinct ?zs" by simp - show "set ?zs \ {(x, y). (R OO S) x y}" using assms(3-4) set_wpull_subset[OF eq] by fastforce - show "remdups (map fst ?zs) = remdups (map fst xs)" unfolding remdups_map_remdups using eq - by(induction xs ys rule: wpull_induct)(auto simp add: set_fst_wpull intro: rev_image_eqI) - show "remdups (map snd ?zs) = remdups (map snd ys)" unfolding remdups_map_remdups using eq - by(induction xs ys rule: wpull_induct)(auto simp add: set_snd_wpull intro: rev_image_eqI) -qed +qualified lemma double_Cons_same [simp]: "double xs ys \ double (x # xs) (x # ys)" + by(auto simp add: double.simps Cons_eq_append_conv) -qualified lift_definition set :: "'a dlist \ 'a set" is List.set . - -qualified lemma map_transfer [transfer_rule]: - "(rel_fun (=) (rel_fun (pcr_dlist (=)) (pcr_dlist (=)))) (\f x. remdups (List.map f x)) Dlist.map" -by(simp add: rel_fun_def dlist.pcr_cr_eq cr_dlist_def Dlist.map_def remdups_remdups) +qualified lemma doubles_Cons_same: "double\<^sup>*\<^sup>* xs ys \ double\<^sup>*\<^sup>* (x # xs) (x # ys)" + by(induction rule: rtranclp_induct)(auto intro: rtranclp.rtrancl_into_rtrancl) -bnf "'a dlist" - map: Dlist.map - sets: set - bd: natLeq - wits: Dlist.empty -unfolding OO_Grp_alt mem_Collect_eq -subgoal by(rule ext)(simp add: dlist_eq_iff) -subgoal by(rule ext)(simp add: dlist_eq_iff remdups_map_remdups) -subgoal by(simp add: dlist_eq_iff set_def cong: list.map_cong) -subgoal by(simp add: set_def fun_eq_iff) -subgoal by(simp add: natLeq_card_order) -subgoal by(simp add: natLeq_cinfinite) -subgoal by(rule ordLess_imp_ordLeq)(simp add: finite_iff_ordLess_natLeq[symmetric] set_def) -subgoal by(rule predicate2I)(transfer; auto simp add: wpull) -subgoal by(simp add: set_def) -done +qualified lemma remdups_into_doubles: "double\<^sup>*\<^sup>* (remdups xs) xs" + by(induction xs)(auto intro: doubles_Cons_same rtranclp.rtrancl_into_rtrancl) + +qualified lemma dlist_eq_into_doubles: "Dlist.dlist_eq \ equivclp double" + by(auto 4 4 simp add: Dlist.dlist_eq_def vimage2p_def + intro: equivclp_trans converse_rtranclp_into_equivclp rtranclp_into_equivclp remdups_into_doubles) + +qualified lemma factor_double_map: "double (map f xs) ys \ \zs. Dlist.dlist_eq xs zs \ ys = map f zs" + by(auto simp add: double.simps Dlist.dlist_eq_def vimage2p_def map_eq_append_conv) + (metis list.simps(9) map_append remdups.simps(2) remdups_append2) + +qualified lemma dlist_eq_set_eq: "Dlist.dlist_eq xs ys \ set xs = set ys" + by(simp add: Dlist.dlist_eq_def vimage2p_def)(metis set_remdups) + +qualified lemma dlist_eq_map_respect: "Dlist.dlist_eq xs ys \ Dlist.dlist_eq (map f xs) (map f ys)" + by(clarsimp simp add: Dlist.dlist_eq_def vimage2p_def)(metis remdups_map_remdups) + +qualified lemma confluent_quotient_dlist: + "confluent_quotient double Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq + (map fst) (map snd) (map fst) (map snd) list_all2 list_all2 list_all2 set set" + by(unfold_locales)(auto intro: strong_confluentp_imp_confluentp strong_confluentp_double + dest: factor_double_map dlist_eq_into_doubles[THEN predicate2D] dlist_eq_set_eq + simp add: list.in_rel list.rel_compp dlist_eq_map_respect Dlist.equivp_dlist_eq equivp_imp_transp) + lifting_update dlist.lifting lifting_forget dlist.lifting end +context begin +interpretation Quotient_dlist: Quotient_dlist . + +lift_bnf (plugins del: code) 'a dlist + subgoal for A B by(rule confluent_quotient.subdistributivity[OF Dlist.confluent_quotient_dlist]) + subgoal by(force dest: Dlist.dlist_eq_set_eq intro: equivp_reflp[OF Dlist.equivp_dlist_eq]) + done + +qualified lemma list_of_dlist_transfer[transfer_rule]: + "bi_unique R \ (rel_fun (Quotient_dlist.pcr_dlist R) (list_all2 R)) remdups list_of_dlist" + unfolding rel_fun_def Quotient_dlist.pcr_dlist_def qcr_dlist_def Dlist.abs_dlist_def + by (auto simp: Abs_dlist_inverse intro!: remdups_transfer[THEN rel_funD]) + +lemma list_of_dlist_map_dlist[simp]: + "list_of_dlist (map_dlist f xs) = remdups (map f (list_of_dlist xs))" + by transfer (auto simp: remdups_map_remdups) + end + + +end diff --git a/src/HOL/Quotient.thy b/src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy +++ b/src/HOL/Quotient.thy @@ -1,859 +1,868 @@ (* Title: HOL/Quotient.thy Author: Cezary Kaliszyk and Christian Urban *) section \Definition of Quotient Types\ theory Quotient imports Lifting keywords "print_quotmapsQ3" "print_quotientsQ3" "print_quotconsts" :: diag and "quotient_type" :: thy_goal_defn and "/" and "quotient_definition" :: thy_goal_defn and "copy_bnf" :: thy_defn and "lift_bnf" :: thy_goal_defn begin text \ Basic definition for equivalence relations that are represented by predicates. \ text \Composition of Relations\ abbreviation rel_conj :: "('a \ 'b \ bool) \ ('b \ 'a \ bool) \ 'a \ 'b \ bool" (infixr "OOO" 75) where "r1 OOO r2 \ r1 OO r2 OO r1" lemma eq_comp_r: shows "((=) OOO R) = R" by (auto simp add: fun_eq_iff) context includes lifting_syntax begin subsection \Quotient Predicate\ definition "Quotient3 R Abs Rep \ (\a. Abs (Rep a) = a) \ (\a. R (Rep a) (Rep a)) \ (\r s. R r s \ R r r \ R s s \ Abs r = Abs s)" lemma Quotient3I: assumes "\a. Abs (Rep a) = a" and "\a. R (Rep a) (Rep a)" and "\r s. R r s \ R r r \ R s s \ Abs r = Abs s" shows "Quotient3 R Abs Rep" using assms unfolding Quotient3_def by blast context fixes R Abs Rep assumes a: "Quotient3 R Abs Rep" begin lemma Quotient3_abs_rep: "Abs (Rep a) = a" using a unfolding Quotient3_def by simp lemma Quotient3_rep_reflp: "R (Rep a) (Rep a)" using a unfolding Quotient3_def by blast lemma Quotient3_rel: "R r r \ R s s \ Abs r = Abs s \ R r s" \ \orientation does not loop on rewriting\ using a unfolding Quotient3_def by blast lemma Quotient3_refl1: "R r s \ R r r" using a unfolding Quotient3_def by fast lemma Quotient3_refl2: "R r s \ R s s" using a unfolding Quotient3_def by fast lemma Quotient3_rel_rep: "R (Rep a) (Rep b) \ a = b" using a unfolding Quotient3_def by metis lemma Quotient3_rep_abs: "R r r \ R (Rep (Abs r)) r" using a unfolding Quotient3_def by blast lemma Quotient3_rel_abs: "R r s \ Abs r = Abs s" using a unfolding Quotient3_def by blast lemma Quotient3_symp: "symp R" using a unfolding Quotient3_def using sympI by metis lemma Quotient3_transp: "transp R" using a unfolding Quotient3_def using transpI by (metis (full_types)) lemma Quotient3_part_equivp: "part_equivp R" by (metis Quotient3_rep_reflp Quotient3_symp Quotient3_transp part_equivpI) lemma abs_o_rep: "Abs \ Rep = id" unfolding fun_eq_iff by (simp add: Quotient3_abs_rep) lemma equals_rsp: assumes b: "R xa xb" "R ya yb" shows "R xa ya = R xb yb" using b Quotient3_symp Quotient3_transp by (blast elim: sympE transpE) lemma rep_abs_rsp: assumes b: "R x1 x2" shows "R x1 (Rep (Abs x2))" using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp by metis lemma rep_abs_rsp_left: assumes b: "R x1 x2" shows "R (Rep (Abs x1)) x2" using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp by metis end lemma identity_quotient3: "Quotient3 (=) id id" unfolding Quotient3_def id_def by blast lemma fun_quotient3: assumes q1: "Quotient3 R1 abs1 rep1" and q2: "Quotient3 R2 abs2 rep2" shows "Quotient3 (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)" proof - have "(rep1 ---> abs2) ((abs1 ---> rep2) a) = a" for a using q1 q2 by (simp add: Quotient3_def fun_eq_iff) moreover have "(R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" for a by (rule rel_funI) (insert q1 q2 Quotient3_rel_abs [of R1 abs1 rep1] Quotient3_rel_rep [of R2 abs2 rep2], simp (no_asm) add: Quotient3_def, simp) moreover have "(R1 ===> R2) r s = ((R1 ===> R2) r r \ (R1 ===> R2) s s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s)" for r s proof - have "(R1 ===> R2) r s \ (R1 ===> R2) r r" unfolding rel_fun_def using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] by (metis (full_types) part_equivp_def) moreover have "(R1 ===> R2) r s \ (R1 ===> R2) s s" unfolding rel_fun_def using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] by (metis (full_types) part_equivp_def) moreover have "(R1 ===> R2) r s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s" by (auto simp add: rel_fun_def fun_eq_iff) (use q1 q2 in \unfold Quotient3_def, metis\) moreover have "((R1 ===> R2) r r \ (R1 ===> R2) s s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s) \ (R1 ===> R2) r s" by (auto simp add: rel_fun_def fun_eq_iff) (use q1 q2 in \unfold Quotient3_def, metis map_fun_apply\) ultimately show ?thesis by blast qed ultimately show ?thesis by (intro Quotient3I) (assumption+) qed lemma lambda_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) = (\x. f x)" unfolding fun_eq_iff using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by simp lemma lambda_prs1: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep1 ---> Abs2) (\x. (Abs1 ---> Rep2) f x) = (\x. f x)" unfolding fun_eq_iff using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by simp text\ In the following theorem R1 can be instantiated with anything, but we know some of the types of the Rep and Abs functions; so by solving Quotient assumptions we can get a unique R1 that will be provable; which is why we need to use \apply_rsp\ and not the primed version\ lemma apply_rspQ3: fixes f g::"'a \ 'c" assumes q: "Quotient3 R1 Abs1 Rep1" and a: "(R1 ===> R2) f g" "R1 x y" shows "R2 (f x) (g y)" using a by (auto elim: rel_funE) lemma apply_rspQ3'': assumes "Quotient3 R Abs Rep" and "(R ===> S) f f" shows "S (f (Rep x)) (f (Rep x))" proof - from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient3_rep_reflp) then show ?thesis using assms(2) by (auto intro: apply_rsp') qed subsection \lemmas for regularisation of ball and bex\ lemma ball_reg_eqv: fixes P :: "'a \ bool" assumes a: "equivp R" shows "Ball (Respects R) P = (All P)" using a unfolding equivp_def by (auto simp add: in_respects) lemma bex_reg_eqv: fixes P :: "'a \ bool" assumes a: "equivp R" shows "Bex (Respects R) P = (Ex P)" using a unfolding equivp_def by (auto simp add: in_respects) lemma ball_reg_right: assumes a: "\x. x \ R \ P x \ Q x" shows "All P \ Ball R Q" using a by fast lemma bex_reg_left: assumes a: "\x. x \ R \ Q x \ P x" shows "Bex R Q \ Ex P" using a by fast lemma ball_reg_left: assumes a: "equivp R" shows "(\x. (Q x \ P x)) \ Ball (Respects R) Q \ All P" using a by (metis equivp_reflp in_respects) lemma bex_reg_right: assumes a: "equivp R" shows "(\x. (Q x \ P x)) \ Ex Q \ Bex (Respects R) P" using a by (metis equivp_reflp in_respects) lemma ball_reg_eqv_range: fixes P::"'a \ bool" and x::"'a" assumes a: "equivp R2" shows "(Ball (Respects (R1 ===> R2)) (\f. P (f x)) = All (\f. P (f x)))" apply(rule iffI) apply(rule allI) apply(drule_tac x="\y. f x" in bspec) apply(simp add: in_respects rel_fun_def) apply(rule impI) using a equivp_reflp_symp_transp[of "R2"] apply (auto elim: equivpE reflpE) done lemma bex_reg_eqv_range: assumes a: "equivp R2" shows "(Bex (Respects (R1 ===> R2)) (\f. P (f x)) = Ex (\f. P (f x)))" apply(auto) apply(rule_tac x="\y. f x" in bexI) apply(simp) apply(simp add: Respects_def in_respects rel_fun_def) apply(rule impI) using a equivp_reflp_symp_transp[of "R2"] apply (auto elim: equivpE reflpE) done (* Next four lemmas are unused *) lemma all_reg: assumes a: "\x :: 'a. (P x \ Q x)" and b: "All P" shows "All Q" using a b by fast lemma ex_reg: assumes a: "\x :: 'a. (P x \ Q x)" and b: "Ex P" shows "Ex Q" using a b by fast lemma ball_reg: assumes a: "\x :: 'a. (x \ R \ P x \ Q x)" and b: "Ball R P" shows "Ball R Q" using a b by fast lemma bex_reg: assumes a: "\x :: 'a. (x \ R \ P x \ Q x)" and b: "Bex R P" shows "Bex R Q" using a b by fast lemma ball_all_comm: assumes "\y. (\x\P. A x y) \ (\x. B x y)" shows "(\x\P. \y. A x y) \ (\x. \y. B x y)" using assms by auto lemma bex_ex_comm: assumes "(\y. \x. A x y) \ (\y. \x\P. B x y)" shows "(\x. \y. A x y) \ (\x\P. \y. B x y)" using assms by auto subsection \Bounded abstraction\ definition Babs :: "'a set \ ('a \ 'b) \ 'a \ 'b" where "x \ p \ Babs p m x = m x" lemma babs_rsp: assumes q: "Quotient3 R1 Abs1 Rep1" and a: "(R1 ===> R2) f g" shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" apply (auto simp add: Babs_def in_respects rel_fun_def) apply (subgoal_tac "x \ Respects R1 \ y \ Respects R1") using a apply (simp add: Babs_def rel_fun_def) apply (simp add: in_respects rel_fun_def) using Quotient3_rel[OF q] by metis lemma babs_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f" apply (rule ext) apply (simp add:) apply (subgoal_tac "Rep1 x \ Respects R1") apply (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2]) apply (simp add: in_respects Quotient3_rel_rep[OF q1]) done lemma babs_simp: assumes q: "Quotient3 R1 Abs Rep" shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)" apply(rule iffI) apply(simp_all only: babs_rsp[OF q]) apply(auto simp add: Babs_def rel_fun_def) apply(metis Babs_def in_respects Quotient3_rel[OF q]) done (* If a user proves that a particular functional relation is an equivalence this may be useful in regularising *) lemma babs_reg_eqv: shows "equivp R \ Babs (Respects R) P = P" by (simp add: fun_eq_iff Babs_def in_respects equivp_reflp) (* 3 lemmas needed for proving repabs_inj *) lemma ball_rsp: assumes a: "(R ===> (=)) f g" shows "Ball (Respects R) f = Ball (Respects R) g" using a by (auto simp add: Ball_def in_respects elim: rel_funE) lemma bex_rsp: assumes a: "(R ===> (=)) f g" shows "(Bex (Respects R) f = Bex (Respects R) g)" using a by (auto simp add: Bex_def in_respects elim: rel_funE) lemma bex1_rsp: assumes a: "(R ===> (=)) f g" shows "Ex1 (\x. x \ Respects R \ f x) = Ex1 (\x. x \ Respects R \ g x)" using a by (auto elim: rel_funE simp add: Ex1_def in_respects) (* 2 lemmas needed for cleaning of quantifiers *) lemma all_prs: assumes a: "Quotient3 R absf repf" shows "Ball (Respects R) ((absf ---> id) f) = All f" using a unfolding Quotient3_def Ball_def in_respects id_apply comp_def map_fun_def by metis lemma ex_prs: assumes a: "Quotient3 R absf repf" shows "Bex (Respects R) ((absf ---> id) f) = Ex f" using a unfolding Quotient3_def Bex_def in_respects id_apply comp_def map_fun_def by metis subsection \\Bex1_rel\ quantifier\ definition Bex1_rel :: "('a \ 'a \ bool) \ ('a \ bool) \ bool" where "Bex1_rel R P \ (\x \ Respects R. P x) \ (\x \ Respects R. \y \ Respects R. ((P x \ P y) \ (R x y)))" lemma bex1_rel_aux: "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R x\ \ Bex1_rel R y" unfolding Bex1_rel_def by (metis in_respects) lemma bex1_rel_aux2: "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R y\ \ Bex1_rel R x" unfolding Bex1_rel_def by (metis in_respects) lemma bex1_rel_rsp: assumes a: "Quotient3 R absf repf" shows "((R ===> (=)) ===> (=)) (Bex1_rel R) (Bex1_rel R)" unfolding rel_fun_def by (metis bex1_rel_aux bex1_rel_aux2) lemma ex1_prs: assumes "Quotient3 R absf repf" shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f" apply (auto simp: Bex1_rel_def Respects_def) apply (metis Quotient3_def assms) apply (metis (full_types) Quotient3_def assms) by (meson Quotient3_rel assms) lemma bex1_bexeq_reg: shows "(\!x\Respects R. P x) \ (Bex1_rel R (\x. P x))" by (auto simp add: Ex1_def Bex1_rel_def Bex_def Ball_def in_respects) lemma bex1_bexeq_reg_eqv: assumes a: "equivp R" shows "(\!x. P x) \ Bex1_rel R P" using equivp_reflp[OF a] by (metis (full_types) Bex1_rel_def in_respects) subsection \Various respects and preserve lemmas\ lemma quot_rel_rsp: assumes a: "Quotient3 R Abs Rep" shows "(R ===> R ===> (=)) R R" apply(rule rel_funI)+ by (meson assms equals_rsp) lemma o_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" and q3: "Quotient3 R3 Abs3 Rep3" shows "((Abs2 ---> Rep3) ---> (Abs1 ---> Rep2) ---> (Rep1 ---> Abs3)) (\) = (\)" and "(id ---> (Abs1 ---> id) ---> Rep1 ---> id) (\) = (\)" using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] Quotient3_abs_rep[OF q3] by (simp_all add: fun_eq_iff) lemma o_rsp: "((R2 ===> R3) ===> (R1 ===> R2) ===> (R1 ===> R3)) (\) (\)" "((=) ===> (R1 ===> (=)) ===> R1 ===> (=)) (\) (\)" by (force elim: rel_funE)+ lemma cond_prs: assumes a: "Quotient3 R absf repf" shows "absf (if a then repf b else repf c) = (if a then b else c)" using a unfolding Quotient3_def by auto lemma if_prs: assumes q: "Quotient3 R Abs Rep" shows "(id ---> Rep ---> Rep ---> Abs) If = If" using Quotient3_abs_rep[OF q] by (auto simp add: fun_eq_iff) lemma if_rsp: assumes q: "Quotient3 R Abs Rep" shows "((=) ===> R ===> R ===> R) If If" by force lemma let_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep2 ---> (Abs2 ---> Rep1) ---> Abs1) Let = Let" using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by (auto simp add: fun_eq_iff) lemma let_rsp: shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let" by (force elim: rel_funE) lemma id_rsp: shows "(R ===> R) id id" by auto lemma id_prs: assumes a: "Quotient3 R Abs Rep" shows "(Rep ---> Abs) id = id" by (simp add: fun_eq_iff Quotient3_abs_rep [OF a]) end locale quot_type = fixes R :: "'a \ 'a \ bool" and Abs :: "'a set \ 'b" and Rep :: "'b \ 'a set" assumes equivp: "part_equivp R" and rep_prop: "\y. \x. R x x \ Rep y = Collect (R x)" and rep_inverse: "\x. Abs (Rep x) = x" and abs_inverse: "\c. (\x. ((R x x) \ (c = Collect (R x)))) \ (Rep (Abs c)) = c" and rep_inject: "\x y. (Rep x = Rep y) = (x = y)" begin definition abs :: "'a \ 'b" where "abs x = Abs (Collect (R x))" definition rep :: "'b \ 'a" where "rep a = (SOME x. x \ Rep a)" lemma some_collect: assumes "R r r" shows "R (SOME x. x \ Collect (R r)) = R r" apply simp by (metis assms exE_some equivp[simplified part_equivp_def]) lemma Quotient: shows "Quotient3 R abs rep" unfolding Quotient3_def abs_def rep_def proof (intro conjI allI) fix a r s show x: "R (SOME x. x \ Rep a) (SOME x. x \ Rep a)" proof - obtain x where r: "R x x" and rep: "Rep a = Collect (R x)" using rep_prop[of a] by auto have "R (SOME x. x \ Rep a) x" using r rep some_collect by metis then have "R x (SOME x. x \ Rep a)" using part_equivp_symp[OF equivp] by fast then show "R (SOME x. x \ Rep a) (SOME x. x \ Rep a)" using part_equivp_transp[OF equivp] by (metis \R (SOME x. x \ Rep a) x\) qed have "Collect (R (SOME x. x \ Rep a)) = (Rep a)" by (metis some_collect rep_prop) then show "Abs (Collect (R (SOME x. x \ Rep a))) = a" using rep_inverse by auto have "R r r \ R s s \ Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" proof - assume "R r r" and "R s s" then have "Abs (Collect (R r)) = Abs (Collect (R s)) \ Collect (R r) = Collect (R s)" by (metis abs_inverse) also have "Collect (R r) = Collect (R s) \ (\A x. x \ A) (Collect (R r)) = (\A x. x \ A) (Collect (R s))" by rule simp_all finally show "Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" by simp qed then show "R r s \ R r r \ R s s \ (Abs (Collect (R r)) = Abs (Collect (R s)))" using equivp[simplified part_equivp_def] by metis qed end subsection \Quotient composition\ lemma OOO_quotient3: fixes R1 :: "'a \ 'a \ bool" fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" fixes R2' :: "'a \ 'a \ bool" fixes R2 :: "'b \ 'b \ bool" assumes R1: "Quotient3 R1 Abs1 Rep1" assumes R2: "Quotient3 R2 Abs2 Rep2" assumes Abs1: "\x y. R2' x y \ R1 x x \ R1 y y \ R2 (Abs1 x) (Abs1 y)" assumes Rep1: "\x y. R2 x y \ R2' (Rep1 x) (Rep1 y)" shows "Quotient3 (R1 OO R2' OO R1) (Abs2 \ Abs1) (Rep1 \ Rep2)" proof - have *: "(R1 OOO R2') r r \ (R1 OOO R2') s s \ (Abs2 \ Abs1) r = (Abs2 \ Abs1) s \ (R1 OOO R2') r s" for r s apply safe subgoal for a b c d apply simp apply (rule_tac b="Rep1 (Abs1 r)" in relcomppI) using Quotient3_refl1 R1 rep_abs_rsp apply fastforce apply (rule_tac b="Rep1 (Abs1 s)" in relcomppI) apply (metis (full_types) Rep1 Abs1 Quotient3_rel R2 Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1] Quotient3_rel_abs [OF R1]) by (metis Quotient3_rel R1 rep_abs_rsp_left) subgoal for x y apply (drule Abs1) apply (erule Quotient3_refl2 [OF R1]) apply (erule Quotient3_refl1 [OF R1]) apply (drule Quotient3_refl1 [OF R2], drule Rep1) by (metis (full_types) Quotient3_def R1 relcompp.relcompI) subgoal for x y apply (drule Abs1) apply (erule Quotient3_refl2 [OF R1]) apply (erule Quotient3_refl1 [OF R1]) apply (drule Quotient3_refl2 [OF R2], drule Rep1) by (metis (full_types) Quotient3_def R1 relcompp.relcompI) subgoal for x y by simp (metis (full_types) Abs1 Quotient3_rel R1 R2) done show ?thesis apply (rule Quotient3I) using * apply (simp_all add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1]) apply (metis Quotient3_rep_reflp R1 R2 Rep1 relcompp.relcompI) done qed lemma OOO_eq_quotient3: fixes R1 :: "'a \ 'a \ bool" fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" assumes R1: "Quotient3 R1 Abs1 Rep1" assumes R2: "Quotient3 (=) Abs2 Rep2" shows "Quotient3 (R1 OOO (=)) (Abs2 \ Abs1) (Rep1 \ Rep2)" using assms by (rule OOO_quotient3) auto subsection \Quotient3 to Quotient\ lemma Quotient3_to_Quotient: assumes "Quotient3 R Abs Rep" and "T \ \x y. R x x \ Abs x = y" shows "Quotient R Abs Rep T" using assms unfolding Quotient3_def by (intro QuotientI) blast+ lemma Quotient3_to_Quotient_equivp: assumes q: "Quotient3 R Abs Rep" and T_def: "T \ \x y. Abs x = y" and eR: "equivp R" shows "Quotient R Abs Rep T" proof (intro QuotientI) fix a show "Abs (Rep a) = a" using q by(rule Quotient3_abs_rep) next fix a show "R (Rep a) (Rep a)" using q by(rule Quotient3_rep_reflp) next fix r s show "R r s = (R r r \ R s s \ Abs r = Abs s)" using q by(rule Quotient3_rel[symmetric]) next show "T = (\x y. R x x \ Abs x = y)" using T_def equivp_reflp[OF eR] by simp qed subsection \ML setup\ text \Auxiliary data for the quotient package\ named_theorems quot_equiv "equivalence relation theorems" and quot_respect "respectfulness theorems" and quot_preserve "preservation theorems" and id_simps "identity simp rules for maps" and quot_thm "quotient theorems" ML_file \Tools/Quotient/quotient_info.ML\ declare [[mapQ3 "fun" = (rel_fun, fun_quotient3)]] lemmas [quot_thm] = fun_quotient3 lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp id_rsp lemmas [quot_preserve] = if_prs o_prs let_prs id_prs lemmas [quot_equiv] = identity_equivp text \Lemmas about simplifying id's.\ lemmas [id_simps] = id_def[symmetric] map_fun_id id_apply id_o o_id eq_comp_r vimage_id text \Translation functions for the lifting process.\ ML_file \Tools/Quotient/quotient_term.ML\ text \Definitions of the quotient types.\ ML_file \Tools/Quotient/quotient_type.ML\ text \Definitions for quotient constants.\ ML_file \Tools/Quotient/quotient_def.ML\ text \ An auxiliary constant for recording some information about the lifted theorem in a tactic. \ definition Quot_True :: "'a \ bool" where "Quot_True x \ True" lemma shows QT_all: "Quot_True (All P) \ Quot_True P" and QT_ex: "Quot_True (Ex P) \ Quot_True P" and QT_ex1: "Quot_True (Ex1 P) \ Quot_True P" and QT_lam: "Quot_True (\x. P x) \ (\x. Quot_True (P x))" and QT_ext: "(\x. Quot_True (a x) \ f x = g x) \ (Quot_True a \ f = g)" by (simp_all add: Quot_True_def ext) lemma QT_imp: "Quot_True a \ Quot_True b" by (simp add: Quot_True_def) context includes lifting_syntax begin text \Tactics for proving the lifted theorems\ ML_file \Tools/Quotient/quotient_tacs.ML\ end subsection \Methods / Interface\ method_setup lifting = \Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms))\ \lift theorems to quotient types\ method_setup lifting_setup = \Attrib.thm >> (fn thm => fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm))\ \set up the three goals for the quotient lifting procedure\ method_setup descending = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt []))\ \decend theorems to the raw level\ method_setup descending_setup = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt []))\ \set up the three goals for the decending theorems\ method_setup partiality_descending = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt []))\ \decend theorems to the raw level\ method_setup partiality_descending_setup = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt []))\ \set up the three goals for the decending theorems\ method_setup regularize = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt))\ \prove the regularization goals from the quotient lifting procedure\ method_setup injection = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt))\ \prove the rep/abs injection goals from the quotient lifting procedure\ method_setup cleaning = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt))\ \prove the cleaning goals from the quotient lifting procedure\ attribute_setup quot_lifted = \Scan.succeed Quotient_Tacs.lifted_attrib\ \lift theorems to quotient types\ no_notation rel_conj (infixr "OOO" 75) section \Lifting of BNFs\ lemma sum_insert_Inl_unit: "x \ A \ (\y. x = Inr y \ Inr y \ B) \ x \ insert (Inl ()) B" by (cases x) (simp_all) lemma lift_sum_unit_vimage_commute: "insert (Inl ()) (Inr ` f -` A) = map_sum id f -` insert (Inl ()) (Inr ` A)" by (auto simp: map_sum_def split: sum.splits) lemma insert_Inl_int_map_sum_unit: "insert (Inl ()) A \ range (map_sum id f) \ {}" by (auto simp: map_sum_def split: sum.splits) lemma image_map_sum_unit_subset: "A \ insert (Inl ()) (Inr ` B) \ map_sum id f ` A \ insert (Inl ()) (Inr ` f ` B)" by auto lemma subset_lift_sum_unitD: "A \ insert (Inl ()) (Inr ` B) \ Inr x \ A \ x \ B" unfolding insert_def by auto lemma UNIV_sum_unit_conv: "insert (Inl ()) (range Inr) = UNIV" unfolding UNIV_sum UNIV_unit image_insert image_empty Un_insert_left sup_bot.left_neutral.. lemma subset_vimage_image_subset: "A \ f -` B \ f ` A \ B" by auto lemma relcompp_mem_Grp_neq_bot: "A \ range f \ {} \ (\x y. x \ A \ y \ A) OO (Grp UNIV f)\\ \ bot" unfolding Grp_def relcompp_apply fun_eq_iff by blast lemma comp_projr_Inr: "projr \ Inr = id" by auto lemma in_rel_sum_in_image_projr: "B \ {(x,y). rel_sum ((=) :: unit \ unit \ bool) A x y} \ Inr ` C = fst ` B \ snd ` B = Inr ` D \ map_prod projr projr ` B \ {(x,y). A x y}" by (force simp: projr_def image_iff dest!: spec[of _ "Inl ()"] split: sum.splits) lemma subset_rel_sumI: "B \ {(x,y). A x y} \ rel_sum ((=) :: unit => unit => bool) A (if x \ B then Inr (fst x) else Inl ()) (if x \ B then Inr (snd x) else Inl ())" by auto lemma relcompp_eq_Grp_neq_bot: "(=) OO (Grp UNIV f)\\ \ bot" unfolding Grp_def relcompp_apply fun_eq_iff by blast lemma rel_fun_rel_OO1: "(rel_fun Q (rel_fun R (=))) A B \ conversep Q OO A OO R \ B" by (auto simp: rel_fun_def) lemma rel_fun_rel_OO2: "(rel_fun Q (rel_fun R (=))) A B \ Q OO B OO conversep R \ A" by (auto simp: rel_fun_def) lemma rel_sum_eq2_nonempty: "rel_sum (=) A OO rel_sum (=) B \ bot" by (auto simp: fun_eq_iff relcompp_apply intro!: exI[of _ "Inl _"]) lemma rel_sum_eq3_nonempty: "rel_sum (=) A OO (rel_sum (=) B OO rel_sum (=) C) \ bot" by (auto simp: fun_eq_iff relcompp_apply intro!: exI[of _ "Inl _"]) lemma hypsubst: "A = B \ x \ B \ (x \ A \ P) \ P" by simp lemma Quotient_crel_quotient: "Quotient R Abs Rep T \ equivp R \ T \ (\x y. Abs x = y)" by (drule Quotient_cr_rel) (auto simp: fun_eq_iff equivp_reflp intro!: eq_reflection) lemma Quotient_crel_typedef: "Quotient (eq_onp P) Abs Rep T \ T \ (\x y. x = Rep y)" unfolding Quotient_def by (auto 0 4 simp: fun_eq_iff eq_onp_def intro: sym intro!: eq_reflection) lemma Quotient_crel_typecopy: "Quotient (=) Abs Rep T \ T \ (\x y. x = Rep y)" by (subst (asm) eq_onp_True[symmetric]) (rule Quotient_crel_typedef) lemma equivp_add_relconj: assumes equiv: "equivp R" "equivp R'" and le: "S OO T OO U \ R OO STU OO R'" shows "R OO S OO T OO U OO R' \ R OO STU OO R'" proof - have trans: "R OO R \ R" "R' OO R' \ R'" using equiv unfolding equivp_reflp_symp_transp transp_relcompp by blast+ have "R OO S OO T OO U OO R' = R OO (S OO T OO U) OO R'" unfolding relcompp_assoc .. also have "\ \ R OO (R OO STU OO R') OO R'" by (intro le relcompp_mono order_refl) also have "\ \ (R OO R) OO STU OO (R' OO R')" unfolding relcompp_assoc .. also have "\ \ R OO STU OO R'" by (intro trans relcompp_mono order_refl) finally show ?thesis . qed lemma Grp_conversep_eq_onp: "((BNF_Def.Grp UNIV f)\\ OO BNF_Def.Grp UNIV f) = eq_onp (\x. x \ range f)" by (auto simp: fun_eq_iff Grp_def eq_onp_def image_iff) lemma Grp_conversep_nonempty: "(BNF_Def.Grp UNIV f)\\ OO BNF_Def.Grp UNIV f \ bot" by (auto simp: fun_eq_iff Grp_def) lemma relcomppI2: "r a b \ s b c \ t c d \ (r OO s OO t) a d" by (auto) lemma rel_conj_eq_onp: "equivp R \ rel_conj R (eq_onp P) \ R" by (auto simp: eq_onp_def transp_def equivp_def) lemma Quotient_Quotient3: "Quotient R Abs Rep T \ Quotient3 R Abs Rep" unfolding Quotient_def Quotient3_def by blast -lemma bnf_lift_Quotient_equivp: "Quotient R Abs Rep T \ equivp R \ True" - by auto +lemma Quotient_reflp_imp_equivp: "Quotient R Abs Rep T \ reflp R \ equivp R" + using Quotient_symp Quotient_transp equivpI by blast + +lemma Quotient_eq_onp_typedef: + "Quotient (eq_onp P) Abs Rep cr \ type_definition Rep Abs {x. P x}" + unfolding Quotient_def eq_onp_def + by unfold_locales auto + +lemma Quotient_eq_onp_type_copy: + "Quotient (=) Abs Rep cr \ type_definition Rep Abs UNIV" + unfolding Quotient_def eq_onp_def + by unfold_locales auto ML_file "Tools/BNF/bnf_lift.ML" hide_fact sum_insert_Inl_unit lift_sum_unit_vimage_commute insert_Inl_int_map_sum_unit image_map_sum_unit_subset subset_lift_sum_unitD UNIV_sum_unit_conv subset_vimage_image_subset relcompp_mem_Grp_neq_bot comp_projr_Inr in_rel_sum_in_image_projr subset_rel_sumI relcompp_eq_Grp_neq_bot rel_fun_rel_OO1 rel_fun_rel_OO2 rel_sum_eq2_nonempty rel_sum_eq3_nonempty hypsubst equivp_add_relconj Grp_conversep_eq_onp Grp_conversep_nonempty relcomppI2 rel_conj_eq_onp - bnf_lift_Quotient_equivp Quotient_Quotient3 + Quotient_reflp_imp_equivp Quotient_Quotient3 end - diff --git a/src/HOL/ROOT b/src/HOL/ROOT --- a/src/HOL/ROOT +++ b/src/HOL/ROOT @@ -1,1132 +1,1131 @@ chapter HOL session HOL (main) = Pure + description " Classical Higher-order Logic. " options [strict_facts] sessions Tools theories Main (global) Complex_Main (global) document_files "root.bib" "root.tex" session "HOL-Proofs" (timing) in Proofs = Pure + description " HOL-Main with explicit proof terms. " options [quick_and_dirty = false, record_proofs = 2, parallel_limit = 500] sessions "HOL-Library" theories "HOL-Library.Realizers" session "HOL-Library" (main timing) in Library = HOL + description " Classical Higher-order Logic -- batteries included. " theories Library (*conflicting type class instantiations and dependent applications*) Finite_Lattice List_Lexorder Prefix_Order Product_Lexorder Product_Order Subseq_Order (*conflicting syntax*) Datatype_Records (*data refinements and dependent applications*) AList_Mapping Code_Binary_Nat Code_Prolog Code_Real_Approx_By_Float Code_Target_Numeral DAList DAList_Multiset RBT_Mapping RBT_Set (*printing modifications*) OptionalSugar (*prototypic tools*) Predicate_Compile_Quickcheck (*legacy tools*) Old_Datatype Old_Recdef Realizers Refute document_files "root.bib" "root.tex" session "HOL-Analysis" (main timing) in Analysis = HOL + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Library" "HOL-Computational_Algebra" theories Analysis document_files "root.tex" "root.bib" session "HOL-Complex_Analysis" (main timing) in Complex_Analysis = "HOL-Analysis" + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] theories Complex_Analysis document_files "root.tex" "root.bib" session "HOL-Analysis-ex" in "Analysis/ex" = "HOL-Analysis" + theories Approximations Metric_Arith_Examples session "HOL-Homology" (timing) in Homology = "HOL-Analysis" + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Algebra" theories Homology document_files "root.tex" session "HOL-Computational_Algebra" (main timing) in "Computational_Algebra" = "HOL-Library" + theories Computational_Algebra (*conflicting type class instantiations and dependent applications*) Field_as_Ring session "HOL-Real_Asymp" in Real_Asymp = HOL + sessions "HOL-Decision_Procs" theories Real_Asymp Real_Asymp_Approx Real_Asymp_Examples session "HOL-Real_Asymp-Manual" in "Real_Asymp/Manual" = "HOL-Real_Asymp" + theories Real_Asymp_Doc document_files (in "~~/src/Doc") "iman.sty" "extra.sty" "isar.sty" document_files "root.tex" "style.sty" session "HOL-Hahn_Banach" in Hahn_Banach = "HOL-Library" + description " Author: Gertrud Bauer, TU Munich The Hahn-Banach theorem for real vector spaces. This is the proof of the Hahn-Banach theorem for real vectorspaces, following H. Heuser, Funktionalanalysis, p. 228 -232. The Hahn-Banach theorem is one of the fundamental theorems of functional analysis. It is a conclusion of Zorn's lemma. Two different formaulations of the theorem are presented, one for general real vectorspaces and its application to normed vectorspaces. The theorem says, that every continous linearform, defined on arbitrary subspaces (not only one-dimensional subspaces), can be extended to a continous linearform on the whole vectorspace. " sessions "HOL-Analysis" theories Hahn_Banach document_files "root.bib" "root.tex" session "HOL-Induct" in Induct = "HOL-Library" + description " Examples of (Co)Inductive Definitions. Comb proves the Church-Rosser theorem for combinators (see http://www.cl.cam.ac.uk/ftp/papers/reports/TR396-lcp-generic-automatic-proof-tools.ps.gz). Mutil is the famous Mutilated Chess Board problem (see http://www.cl.cam.ac.uk/ftp/papers/reports/TR394-lcp-mutilated-chess-board.dvi.gz). PropLog proves the completeness of a formalization of propositional logic (see http://www.cl.cam.ac.uk/Research/Reports/TR312-lcp-set-II.ps.gz). Exp demonstrates the use of iterated inductive definitions to reason about mutually recursive relations. " theories [quick_and_dirty] Common_Patterns theories Nested_Datatype QuoDataType QuoNestedDataType Term SList ABexp Infinitely_Branching_Tree Ordinals Sigma_Algebra Comb PropLog Com document_files "root.tex" session "HOL-IMP" (timing) in IMP = "HOL-Library" + options [document_variants = document] theories BExp ASM Finite_Reachable Denotational Compiler2 Poly_Types Sec_Typing Sec_TypingT Def_Init_Big Def_Init_Small Fold Live Live_True Hoare_Examples Hoare_Sound_Complete VCG Hoare_Total VCG_Total_EX VCG_Total_EX2 Collecting1 Collecting_Examples Abs_Int_Tests Abs_Int1_parity Abs_Int1_const Abs_Int3 Procs_Dyn_Vars_Dyn Procs_Stat_Vars_Dyn Procs_Stat_Vars_Stat C_like OO document_files "root.bib" "root.tex" session "HOL-IMPP" in IMPP = HOL + description \ Author: David von Oheimb Copyright 1999 TUM IMPP -- An imperative language with procedures. This is an extension of IMP with local variables and mutually recursive procedures. For documentation see "Hoare Logic for Mutual Recursion and Local Variables" (https://isabelle.in.tum.de/Bali/papers/FSTTCS99.html). \ theories EvenOdd session "HOL-Data_Structures" (timing) in Data_Structures = HOL + options [document_variants = document] sessions "HOL-Number_Theory" theories [document = false] Less_False theories Sorting Balance Tree_Map Interval_Tree AVL_Map RBT_Set2 RBT_Map Tree23_Map Tree234_Map Brother12_Map AA_Map Set2_Join_RBT Array_Braun Trie_Fun Trie_Map Tries_Binary Leftist_Heap Binomial_Heap document_files "root.tex" "root.bib" session "HOL-Import" in Import = HOL + theories HOL_Light_Maps theories [condition = HOL_LIGHT_BUNDLE] HOL_Light_Import session "HOL-Number_Theory" (main timing) in Number_Theory = "HOL-Computational_Algebra" + description " Fundamental Theorem of Arithmetic, Chinese Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, some lemmas for Quadratic Reciprocity. " sessions "HOL-Algebra" theories Number_Theory document_files "root.tex" session "HOL-Hoare" in Hoare = HOL + description " Verification of imperative programs (verification conditions are generated automatically from pre/post conditions and loop invariants). " theories Hoare document_files "root.bib" "root.tex" session "HOL-Hoare_Parallel" (timing) in Hoare_Parallel = HOL + description " Verification of shared-variable imperative programs a la Owicki-Gries. (verification conditions are generated automatically). " theories Hoare_Parallel document_files "root.bib" "root.tex" session "HOL-Codegenerator_Test" in Codegenerator_Test = "HOL-Number_Theory" + sessions "HOL-Data_Structures" "HOL-ex" theories Generate Generate_Binary_Nat Generate_Target_Nat Generate_Efficient_Datastructures Code_Lazy_Test Code_Test_PolyML Code_Test_Scala theories [condition = ISABELLE_GHC] Code_Test_GHC theories [condition = ISABELLE_MLTON] Code_Test_MLton theories [condition = ISABELLE_OCAMLFIND] Code_Test_OCaml theories [condition = ISABELLE_SMLNJ] Code_Test_SMLNJ session "HOL-Metis_Examples" (timing) in Metis_Examples = "HOL-Library" + description " Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Testing Metis and Sledgehammer. " sessions "HOL-Decision_Procs" theories Abstraction Big_O Binary_Tree Clausification Message Proxies Tarski Trans_Closure Sets session "HOL-Nitpick_Examples" in Nitpick_Examples = "HOL-Library" + description " Author: Jasmin Blanchette, TU Muenchen Copyright 2009 " theories [quick_and_dirty] Nitpick_Examples session "HOL-Algebra" (main timing) in Algebra = "HOL-Computational_Algebra" + description " Author: Clemens Ballarin, started 24 September 1999, and many others The Isabelle Algebraic Library. " sessions "HOL-Cardinals" theories (* Orders and Lattices *) Galois_Connection (* Knaster-Tarski theorem and Galois connections *) (* Groups *) FiniteProduct (* Product operator for commutative groups *) Sylow (* Sylow's theorem *) Bij (* Automorphism Groups *) Multiplicative_Group Zassenhaus (* The Zassenhaus lemma *) (* Rings *) Divisibility (* Rings *) IntRing (* Ideals and residue classes *) UnivPoly (* Polynomials *) (* Main theory *) Algebra document_files "root.bib" "root.tex" session "HOL-Auth" (timing) in Auth = "HOL-Library" + description " A new approach to verifying authentication protocols. " directories "Smartcard" "Guard" theories Auth_Shared Auth_Public "Smartcard/Auth_Smartcard" "Guard/Auth_Guard_Shared" "Guard/Auth_Guard_Public" document_files "root.tex" session "HOL-UNITY" (timing) in UNITY = "HOL-Auth" + description " Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge Verifying security protocols using Chandy and Misra's UNITY formalism. " directories "Simple" "Comp" theories (*Basic meta-theory*) UNITY_Main (*Simple examples: no composition*) "Simple/Deadlock" "Simple/Common" "Simple/Network" "Simple/Token" "Simple/Channel" "Simple/Lift" "Simple/Mutex" "Simple/Reach" "Simple/Reachability" (*Verifying security protocols using UNITY*) "Simple/NSP_Bad" (*Example of composition*) "Comp/Handshake" (*Universal properties examples*) "Comp/Counter" "Comp/Counterc" "Comp/Priority" "Comp/TimerArray" "Comp/Progress" "Comp/Alloc" "Comp/AllocImpl" "Comp/Client" (*obsolete*) ELT document_files "root.tex" session "HOL-Unix" in Unix = "HOL-Library" + options [print_mode = "no_brackets,no_type_brackets"] theories Unix document_files "root.bib" "root.tex" session "HOL-ZF" in ZF = "HOL-Library" + theories MainZF Games document_files "root.tex" session "HOL-Imperative_HOL" (timing) in Imperative_HOL = "HOL-Library" + options [print_mode = "iff,no_brackets"] directories "ex" theories Imperative_HOL_ex document_files "root.bib" "root.tex" session "HOL-Decision_Procs" (timing) in Decision_Procs = "HOL-Algebra" + description " Various decision procedures, typically involving reflection. " directories "ex" theories Decision_Procs session "HOL-Proofs-ex" in "Proofs/ex" = "HOL-Proofs" + sessions "HOL-Isar_Examples" theories Hilbert_Classical Proof_Terms XML_Data session "HOL-Proofs-Extraction" (timing) in "Proofs/Extraction" = "HOL-Proofs" + description " Examples for program extraction in Higher-Order Logic. " options [quick_and_dirty = false] sessions "HOL-Computational_Algebra" theories Greatest_Common_Divisor Warshall Higman_Extraction Pigeonhole Euclid document_files "root.bib" "root.tex" session "HOL-Proofs-Lambda" (timing) in "Proofs/Lambda" = "HOL-Proofs" + description \ Lambda Calculus in de Bruijn's Notation. This session defines lambda-calculus terms with de Bruijn indixes and proves confluence of beta, eta and beta+eta. The paper "More Church-Rosser Proofs (in Isabelle/HOL)" describes the whole theory (see http://www.in.tum.de/~nipkow/pubs/jar2001.html). \ options [print_mode = "no_brackets", quick_and_dirty = false] sessions "HOL-Library" theories Eta StrongNorm Standardization WeakNorm document_files "root.bib" "root.tex" session "HOL-Prolog" in Prolog = HOL + description " Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) A bare-bones implementation of Lambda-Prolog. This is a simple exploratory implementation of Lambda-Prolog in HOL, including some minimal examples (in Test.thy) and a more typical example of a little functional language and its type system. " theories Test Type session "HOL-MicroJava" (timing) in MicroJava = "HOL-Library" + description " Formalization of a fragment of Java, together with a corresponding virtual machine and a specification of its bytecode verifier and a lightweight bytecode verifier, including proofs of type-safety. " sessions "HOL-Eisbach" directories "BV" "Comp" "DFA" "J" "JVM" theories MicroJava document_files "introduction.tex" "root.bib" "root.tex" session "HOL-NanoJava" in NanoJava = HOL + description " Hoare Logic for a tiny fragment of Java. " theories Example document_files "root.bib" "root.tex" session "HOL-Bali" (timing) in Bali = "HOL-Library" + theories AxExample AxSound AxCompl Trans TypeSafe document_files "root.tex" session "HOL-IOA" in IOA = HOL + description \ Author: Tobias Nipkow and Konrad Slind and Olaf Müller Copyright 1994--1996 TU Muenchen The meta-theory of I/O-Automata in HOL. This formalization has been significantly changed and extended, see HOLCF/IOA. There are also the proofs of two communication protocols which formerly have been here. @inproceedings{Nipkow-Slind-IOA, author={Tobias Nipkow and Konrad Slind}, title={{I/O} Automata in {Isabelle/HOL}}, booktitle={Proc.\ TYPES Workshop 1994}, publisher=Springer, series=LNCS, note={To appear}} ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/ioa.ps.gz and @inproceedings{Mueller-Nipkow, author={Olaf M\"uller and Tobias Nipkow}, title={Combining Model Checking and Deduction for {I/O}-Automata}, booktitle={Proc.\ TACAS Workshop}, organization={Aarhus University, BRICS report}, year=1995} ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/tacas.dvi.gz \ theories Solve session "HOL-Lattice" in Lattice = HOL + description " Author: Markus Wenzel, TU Muenchen Basic theory of lattices and orders. " theories CompleteLattice document_files "root.tex" session "HOL-ex" (timing) in ex = "HOL-Number_Theory" + description " Miscellaneous examples for Higher-Order Logic. " theories Adhoc_Overloading_Examples Antiquote Argo_Examples Arith_Examples Ballot BinEx Birthday_Paradox Bit_Lists Bubblesort CTL Cartouche_Examples Case_Product Chinese Classical Code_Binary_Nat_examples Code_Lazy_Demo Code_Timing Coercion_Examples Coherent Commands Computations Conditional_Parametricity_Examples Cubic_Quartic Datatype_Record_Examples Dedekind_Real Erdoes_Szekeres Eval_Examples Executable_Relation Execute_Choice Functions Function_Growth Gauge_Integration Groebner_Examples Guess HarmonicSeries Hebrew Hex_Bin_Examples IArray_Examples Iff_Oracle Induction_Schema Intuitionistic Join_Theory Lagrange List_to_Set_Comprehension_Examples LocaleTest2 "ML" MergeSort MonoidGroup Multiquote NatSum Normalization_by_Evaluation PER Parallel_Example Peano_Axioms Perm_Fragments PresburgerEx Primrec Pythagoras Quicksort Radix_Sort Records Reflection_Examples Refute_Examples Residue_Ring Rewrite_Examples SOS SOS_Cert Seq Serbian Set_Comprehension_Pointfree_Examples Set_Theory Simproc_Tests Simps_Case_Conv_Examples Sketch_and_Explore Sorting_Algorithms_Examples Sqrt Sqrt_Script Sudoku Sum_of_Powers Tarski Termination ThreeDivides Transfer_Debug Transfer_Int_Nat Transitive_Closure_Table_Ex Tree23 Triangular_Numbers Unification While_Combinator_Example Word veriT_Preprocessing theories [skip_proofs = false] SAT_Examples Meson_Test session "HOL-Isar_Examples" in Isar_Examples = "HOL-Computational_Algebra" + description " Miscellaneous Isabelle/Isar examples. " options [quick_and_dirty] theories Knaster_Tarski Peirce Drinker Cantor Structured_Statements Basic_Logic Expr_Compiler Fibonacci Group Group_Context Group_Notepad Hoare_Ex Mutilated_Checkerboard Puzzle Summation First_Order_Logic Higher_Order_Logic document_files "root.bib" "root.tex" session "HOL-Eisbach" in Eisbach = HOL + description \ The Eisbach proof method language and "match" method. \ sessions FOL "HOL-Analysis" theories Eisbach Tests Examples Examples_FOL Example_Metric session "HOL-SET_Protocol" (timing) in SET_Protocol = "HOL-Library" + description " Verification of the SET Protocol. " theories SET_Protocol document_files "root.tex" session "HOL-Matrix_LP" in Matrix_LP = "HOL-Library" + description " Two-dimensional matrices and linear programming. " directories "Compute_Oracle" theories Cplex document_files "root.tex" session "HOL-TLA" in TLA = HOL + description " Lamport's Temporal Logic of Actions. " theories TLA session "HOL-TLA-Inc" in "TLA/Inc" = "HOL-TLA" + theories Inc session "HOL-TLA-Buffer" in "TLA/Buffer" = "HOL-TLA" + theories DBuffer session "HOL-TLA-Memory" in "TLA/Memory" = "HOL-TLA" + theories MemoryImplementation session "HOL-TPTP" in TPTP = "HOL-Library" + description " Author: Jasmin Blanchette, TU Muenchen Author: Nik Sultana, University of Cambridge Copyright 2011 TPTP-related extensions. " theories ATP_Theory_Export MaSh_Eval TPTP_Interpret THF_Arith TPTP_Proof_Reconstruction theories ATP_Problem_Import session "HOL-Probability" (main timing) in "Probability" = "HOL-Analysis" + theories Probability document_files "root.tex" session "HOL-Probability-ex" (timing) in "Probability/ex" = "HOL-Probability" + theories Dining_Cryptographers Koepf_Duermuth_Countermeasure Measure_Not_CCC session "HOL-Nominal" in Nominal = "HOL-Library" + theories Nominal session "HOL-Nominal-Examples" (timing) in "Nominal/Examples" = "HOL-Nominal" + theories Class3 CK_Machine Compile Contexts Crary CR_Takahashi CR Fsub Height Lambda_mu Lam_Funs LocalWeakening Pattern SN SOS Standardization Support Type_Preservation Weakening W theories [quick_and_dirty] VC_Condition session "HOL-Cardinals" (timing) in Cardinals = HOL + description " Ordinals and Cardinals, Full Theories. " theories Cardinals Bounded_Set document_files "intro.tex" "root.tex" "root.bib" session "HOL-Datatype_Examples" (timing) in Datatype_Examples = "HOL-Library" + description " (Co)datatype Examples. " directories "Derivation_Trees" theories Compat Lambda_Term Process TreeFsetI "Derivation_Trees/Gram_Lang" "Derivation_Trees/Parallel_Composition" Koenig Lift_BNF Milner_Tofte Stream_Processor Cyclic_List - Dlist Free_Idempotent_Monoid LDL TLList Misc_Codatatype Misc_Datatype Misc_Primcorec Misc_Primrec session "HOL-Corec_Examples" (timing) in Corec_Examples = "HOL-Library" + description " Corecursion Examples. " directories "Tests" theories LFilter Paper_Examples Stream_Processor "Tests/Simple_Nesting" "Tests/Iterate_GPV" theories [quick_and_dirty] "Tests/GPV_Bare_Bones" "Tests/Merge_D" "Tests/Merge_Poly" "Tests/Misc_Mono" "Tests/Misc_Poly" "Tests/Small_Concrete" "Tests/Stream_Friends" "Tests/TLList_Friends" "Tests/Type_Class" session "HOL-Word" (main timing) in Word = HOL + sessions "HOL-Library" theories Word More_Word Word_Examples document_files "root.bib" "root.tex" session "HOL-Statespace" in Statespace = HOL + theories [skip_proofs = false] StateSpaceEx document_files "root.tex" session "HOL-Nonstandard_Analysis" (timing) in Nonstandard_Analysis = "HOL-Computational_Algebra" + description " Nonstandard analysis. " theories Nonstandard_Analysis document_files "root.tex" session "HOL-Nonstandard_Analysis-Examples" (timing) in "Nonstandard_Analysis/Examples" = "HOL-Nonstandard_Analysis" + theories NSPrimes session "HOL-Mirabelle" in Mirabelle = HOL + theories Mirabelle_Test session "HOL-Mirabelle-ex" in "Mirabelle/ex" = "HOL-Mirabelle" + options [timeout = 60] theories Ex session "HOL-Word-SMT_Examples" (timing) in SMT_Examples = "HOL-Word" + options [quick_and_dirty] theories Boogie SMT_Examples SMT_Word_Examples SMT_Tests session "HOL-SPARK" in "SPARK" = "HOL-Word" + theories SPARK session "HOL-SPARK-Examples" in "SPARK/Examples" = "HOL-SPARK" + directories "Gcd" "Liseq" "RIPEMD-160" "Sqrt" theories "Gcd/Greatest_Common_Divisor" "Liseq/Longest_Increasing_Subsequence" "RIPEMD-160/F" "RIPEMD-160/Hash" "RIPEMD-160/K_L" "RIPEMD-160/K_R" "RIPEMD-160/R_L" "RIPEMD-160/Round" "RIPEMD-160/R_R" "RIPEMD-160/S_L" "RIPEMD-160/S_R" "Sqrt/Sqrt" export_files (in ".") "*:**.prv" session "HOL-SPARK-Manual" in "SPARK/Manual" = "HOL-SPARK" + options [show_question_marks = false] sessions "HOL-SPARK-Examples" theories Example_Verification VC_Principles Reference Complex_Types document_files "complex_types.ads" "complex_types_app.adb" "complex_types_app.ads" "Gcd.adb" "Gcd.ads" "intro.tex" "loop_invariant.adb" "loop_invariant.ads" "root.bib" "root.tex" "Simple_Gcd.adb" "Simple_Gcd.ads" session "HOL-Mutabelle" in Mutabelle = "HOL-Library" + theories MutabelleExtra session "HOL-Quickcheck_Examples" (timing) in Quickcheck_Examples = "HOL-Library" + theories Quickcheck_Examples Quickcheck_Lattice_Examples Completeness Quickcheck_Interfaces Quickcheck_Nesting_Example theories [condition = ISABELLE_GHC] Hotel_Example Quickcheck_Narrowing_Examples session "HOL-Quotient_Examples" (timing) in Quotient_Examples = "HOL-Algebra" + description " Author: Cezary Kaliszyk and Christian Urban " theories DList Quotient_FSet Quotient_Int Quotient_Message Lift_FSet Lift_Set Lift_Fun Quotient_Rat Lift_DList Int_Pow Lifting_Code_Dt_Test session "HOL-Predicate_Compile_Examples" (timing) in Predicate_Compile_Examples = "HOL-Library" + theories Examples Predicate_Compile_Tests Predicate_Compile_Quickcheck_Examples Specialisation_Examples IMP_1 IMP_2 (* FIXME since 21-Jul-2011 Hotel_Example_Small_Generator *) IMP_3 IMP_4 theories [condition = ISABELLE_SWIPL] Code_Prolog_Examples Context_Free_Grammar_Example Hotel_Example_Prolog Lambda_Example List_Examples theories [condition = ISABELLE_SWIPL, quick_and_dirty] Reg_Exp_Example session "HOL-Types_To_Sets" in Types_To_Sets = HOL + description " Experimental extension of Higher-Order Logic to allow translation of types to sets. " directories "Examples" theories Types_To_Sets "Examples/Prerequisites" "Examples/Finite" "Examples/T2_Spaces" "Examples/Unoverload_Def" "Examples/Linear_Algebra_On" session HOLCF (main timing) in HOLCF = HOL + description " Author: Franz Regensburger Author: Brian Huffman HOLCF -- a semantic extension of HOL by the LCF logic. " sessions "HOL-Library" theories HOLCF (global) document_files "root.tex" session "HOLCF-Tutorial" in "HOLCF/Tutorial" = HOLCF + theories Domain_ex Fixrec_ex New_Domain document_files "root.tex" session "HOLCF-Library" in "HOLCF/Library" = HOLCF + theories HOLCF_Library HOL_Cpo session "HOLCF-IMP" in "HOLCF/IMP" = HOLCF + description " IMP -- A WHILE-language and its Semantics. This is the HOLCF-based denotational semantics of a simple WHILE-language. " sessions "HOL-IMP" theories HoareEx document_files "isaverbatimwrite.sty" "root.tex" "root.bib" session "HOLCF-ex" in "HOLCF/ex" = "HOLCF-Library" + description " Miscellaneous examples for HOLCF. " theories Dnat Dagstuhl Focus_ex Fix2 Hoare Concurrency_Monad Loop Powerdomain_ex Domain_Proofs Letrec Pattern_Match session "HOLCF-FOCUS" in "HOLCF/FOCUS" = "HOLCF-Library" + description \ FOCUS: a theory of stream-processing functions Isabelle/HOLCF. For introductions to FOCUS, see "The Design of Distributed Systems - An Introduction to FOCUS" http://www4.in.tum.de/publ/html.php?e=2 "Specification and Refinement of a Buffer of Length One" http://www4.in.tum.de/publ/html.php?e=15 "Specification and Development of Interactive Systems: Focus on Streams, Interfaces, and Refinement" http://www4.in.tum.de/publ/html.php?e=321 \ theories Fstreams FOCUS Buffer_adm session IOA (timing) in "HOLCF/IOA" = HOLCF + description " Author: Olaf Mueller Copyright 1997 TU München A formalization of I/O automata in HOLCF. The distribution contains simulation relations, temporal logic, and an abstraction theory. Everything is based upon a domain-theoretic model of finite and infinite sequences. " theories Abstraction session "IOA-ABP" in "HOLCF/IOA/ABP" = IOA + description " Author: Olaf Mueller The Alternating Bit Protocol performed in I/O-Automata. " theories Correctness Spec session "IOA-NTP" in "HOLCF/IOA/NTP" = IOA + description " Author: Tobias Nipkow & Konrad Slind A network transmission protocol, performed in the I/O automata formalization by Olaf Mueller. " theories Correctness session "IOA-Storage" in "HOLCF/IOA/Storage" = IOA + description " Author: Olaf Mueller Memory storage case study. " theories Correctness session "IOA-ex" in "HOLCF/IOA/ex" = IOA + description " Author: Olaf Mueller " theories TrivEx TrivEx2 diff --git a/src/HOL/Tools/BNF/bnf_lift.ML b/src/HOL/Tools/BNF/bnf_lift.ML --- a/src/HOL/Tools/BNF/bnf_lift.ML +++ b/src/HOL/Tools/BNF/bnf_lift.ML @@ -1,2090 +1,2127 @@ (* Title: HOL/Tools/BNF/bnf_lift.ML Author: Julian Biendarra, TU Muenchen Author: Basil Fürer, ETH Zurich Author: Joshua Schneider, ETH Zurich Author: Dmitriy Traytel, ETH Zurich Copyright 2015, 2019 Lifting of BNFs through typedefs. *) signature BNF_LIFT = sig datatype lift_bnf_option = Plugins_Option of Proof.context -> Plugin_Name.filter | No_Warn_Wits | No_Warn_Transfer val copy_bnf: (((lift_bnf_option list * (binding option * (string * sort option)) list) * string) * thm option) * (binding * binding * binding) -> local_theory -> local_theory val copy_bnf_cmd: (((lift_bnf_option list * (binding option * (string * string option)) list) * string) * (Facts.ref * Token.src list) option) * (binding * binding * binding) -> local_theory -> local_theory val lift_bnf: ((((lift_bnf_option list * (binding option * (string * sort option)) list) * - string) * term list option) * thm option) * (binding * binding * binding) -> + string) * term list option) * thm list option) * (binding * binding * binding) -> ({context: Proof.context, prems: thm list} -> tactic) list -> local_theory -> local_theory val lift_bnf_cmd: ((((lift_bnf_option list * (binding option * (string * string option)) list) * - string) * string list) * (Facts.ref * Token.src list) option) * + string) * string list) * (Facts.ref * Token.src list) list option) * (binding * binding * binding) -> local_theory -> Proof.state end structure BNF_Lift : BNF_LIFT = struct open Ctr_Sugar_Tactics open BNF_Util open BNF_Comp open BNF_Def datatype lift_bnf_option = Plugins_Option of Proof.context -> Plugin_Name.filter | No_Warn_Wits | No_Warn_Transfer; -datatype lift_thm = Typedef of thm | Quotient of thm +datatype equiv_thm = Typedef | Quotient of thm (** Util **) fun last2 [x, y] = ([], (x, y)) | last2 (x :: xs) = last2 xs |>> (fn y => x :: y) | last2 [] = raise Match; fun strip3 thm = (case Term.strip_comb (HOLogic.dest_Trueprop (Thm.prop_of thm)) of (_, [x1, x2, x3]) => (x1, x2, x3) | _ => error "strip3: wrong number of arguments"); val mk_Free = yield_singleton o mk_Frees; fun TWICE t = t THEN' t; fun prove lthy fvars tm tac = Goal.prove_sorry lthy (map (fst o dest_Free) fvars) [] tm (fn {context, ...} => tac context); (** Term construction **) fun mk_relT aT bT = aT --> bT --> HOLogic.boolT; fun mk_relcompp r s = let val (rT, sT) = apply2 fastype_of (r, s); val ((xT, _), (_, zTs)) = apply2 dest_funT (rT, sT); val T = rT --> sT --> mk_relT xT (fst (dest_funT zTs)); in Const (@{const_name relcompp}, T) $ r $ s end; val mk_OO = fold_rev mk_relcompp; (* option from sum *) fun mk_MaybeT T = mk_sumT (HOLogic.unitT, T); fun mk_Nothing T = BNF_FP_Util.mk_Inl T HOLogic.unit; val Just_const = BNF_FP_Util.Inr_const HOLogic.unitT; fun mk_Just tm = Just_const (fastype_of tm) $ tm; fun Maybe_map_const T = let val (xT, yT) = dest_funT T in Const (@{const_name map_sum}, (HOLogic.unitT --> HOLogic.unitT) --> T --> mk_MaybeT xT --> mk_MaybeT yT) $ HOLogic.id_const HOLogic.unitT end; fun mk_Maybe_map tm = Maybe_map_const (fastype_of tm) $ tm; fun fromJust_const T = Const (@{const_name sum.projr}, mk_MaybeT T --> T) fun rel_Maybe_const T U = Const (@{const_name rel_sum}, (HOLogic.unitT --> HOLogic.unitT --> HOLogic.boolT) --> (T --> U --> HOLogic.boolT) --> mk_MaybeT T --> mk_MaybeT U --> HOLogic.boolT) $ HOLogic.eq_const HOLogic.unitT fun set_Maybe_const T = Const (@{const_name Basic_BNFs.setr}, mk_MaybeT T --> HOLogic.mk_setT T) fun Inf_const T = Const (@{const_name Inf}, HOLogic.mk_setT T --> T); fun Image_const T = let val relT = HOLogic.mk_setT (HOLogic.mk_prodT (T, T)); val setT = HOLogic.mk_setT T; in Const (@{const_name Image}, relT --> setT --> setT) end; fun bot_const T = Const (@{const_name bot}, T); fun mk_insert x S = Const (@{const_name Set.insert}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S; fun mk_vimage f s = let val (xT, yT) = dest_funT (fastype_of f) in Const (@{const_name vimage}, (xT --> yT) --> HOLogic.mk_setT yT --> HOLogic.mk_setT xT) $ f $ s end; fun mk_case_prod (x, y) tm = let val ((x, xT), (y, yT)) = apply2 dest_Free (x, y); val prodT = HOLogic.mk_prodT (xT, yT); in HOLogic.Collect_const prodT $ (Const (@{const_name case_prod}, (xT --> yT --> HOLogic.boolT) --> prodT --> HOLogic.boolT) $ absfree (x, xT) (absfree (y, yT) tm)) end; fun mk_Trueprop_implies (ps, c) = Logic.list_implies (map HOLogic.mk_Trueprop ps, HOLogic.mk_Trueprop c); fun mk_Collect (v, tm) = let val (n, T) = dest_Free v in HOLogic.mk_Collect (n, T, tm) end; (** witnesses **) fun prepare_wits is_quotient RepT wits opts alphas wits_F var_as var_as' sets lthy = let fun binder_types_until_eq V T = let fun strip (TU as Type ("fun", [T, U])) = if V = TU then [] else T :: strip U | strip T = if V = T then [] else error ("Bad type for witness: " ^ quote (Syntax.string_of_typ lthy T)); in strip T end; val Iwits = the_default wits_F (Option.map (map (`(map (fn T => find_index (fn U => T = U) alphas) o binder_types_until_eq RepT o fastype_of))) wits); val Iwits = if is_quotient then let val subsumed_Iwits = filter (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) wits_F) Iwits; val _ = if null subsumed_Iwits orelse exists (fn No_Warn_Wits => true | _ => false) opts then () else let val (suffix1, suffix2, be) = (if length subsumed_Iwits = 1 then ("", "", "is") else ("s", "es", "are")) in subsumed_Iwits |> map (Syntax.pretty_typ lthy o fastype_of o snd) |> Pretty.big_list ("The following type" ^ suffix1 ^ " of nonemptiness witness" ^ suffix2 ^ " of the raw type's BNF " ^ be ^ " subsumed by the existing raw witnesses:") |> (fn pt => Pretty.chunks [pt, Pretty.para ("You do not need to lift these subsumed witnesses.")]) |> Pretty.string_of |> warning end; in filter_out (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) wits_F) Iwits end else Iwits; val wit_goals = maps (BNF_Def.mk_wit_goals var_as var_as' sets) Iwits; val lost_wits = if is_quotient then [] else filter_out (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) Iwits) wits_F; val _ = if null lost_wits orelse exists (fn No_Warn_Wits => true | _ => false) opts then () else let val what = (if is_quotient then "quotient type" else "typedef"); val (suffix1, suffix2, be) = (if length lost_wits = 1 then ("", "", "was") else ("s", "es", "were")) in lost_wits |> map (Syntax.pretty_typ lthy o fastype_of o snd) |> Pretty.big_list ("The following type" ^ suffix1 ^ " of nonemptiness witness" ^ suffix2 ^ " of the raw type's BNF " ^ be ^ " lost:") |> (fn pt => Pretty.chunks [pt, Pretty.para ("You can specify a liftable witness (e.g., a term of one of the above types\ \ that satisfies the " ^ what ^ "'s invariant)\ \ using the annotation [wits: ].")]) |> Pretty.string_of |> warning end; in (Iwits, wit_goals) end; (** transfer theorems **) fun mk_transfer_thms' bnf_F bnf_G name consts Tss crel_def pcrel_def defs lthy = let val live = length (#alphas Tss); val (pcrel_tm, crel_tm) = apply2 (Thm.prop_of #> Logic.dest_equals #> fst #> head_of) (pcrel_def, crel_def); val (var_Qs, var_Rs) = fold Variable.declare_typ (#alphas Tss @ #deads Tss) lthy |> mk_Frees "Q" (map2 mk_relT (#alphas Tss) (#betas Tss)) ||>> mk_Frees "R" (map2 mk_relT (#gammas Tss) (#deltas Tss)) |> fst; (* get the "pcrel :: args_raw => rep => abs \ bool" term and instantiate types *) val (args_raw, (rep, abs)) = pcrel_tm |> fastype_of |> binder_types |> last2; val thy = Proof_Context.theory_of lthy; val tyenv_match = Vartab.empty |> Sign.typ_match thy ((rep, #rep Tss)) |> Sign.typ_match thy ((abs, #abs Tss)); val args = map (Envir.subst_type tyenv_match) args_raw; val (pcrel_a, pcrel_b) = Envir.subst_term (tyenv_match, Vartab.empty) pcrel_tm |> `(subst_atomic_types (#alphas Tss @ #betas Tss ~~ #gammas Tss @ #deltas Tss)); (* match "crel :: {?a F} \ a G" with "rep_G :: {a F}" *) val tyenv_match = Vartab.empty |> Sign.typ_match thy (crel_tm |> fastype_of |> binder_types |> hd, #rep Tss); val crel_b = Envir.subst_term (tyenv_match, Vartab.empty) crel_tm |> subst_atomic_types (#alphas Tss ~~ #betas Tss) val crel_d = subst_atomic_types (#betas Tss ~~ #deltas Tss) crel_b; (* instantiate pcrel with Qs and Rs*) val dead_args = map binder_types args |> map (fn [T,U] => if T = U then SOME T else NONE | _ => NONE); val parametrize = let fun go (SOME T :: dTs) tms = HOLogic.eq_const T :: go dTs tms | go (_ :: dTs) (tm :: tms) = tm :: go dTs tms | go (_ :: dTs) tms = go dTs tms | go _ _ = []; in go dead_args end; val pcrel_Qs = list_comb (pcrel_b, parametrize var_Qs); val pcrel_Rs = list_comb (pcrel_a, parametrize var_Rs); (* get the order of the dead variables right *) val Ds0 = maps the_list dead_args; val permute_Ds = (mk_T_of_bnf Ds0 (#betas Tss) bnf_G, nth (binder_types (type_of pcrel_Qs)) 1) |> apply2 (fn Type (_, Ts) => Ts | _ => []) |> op~~ |> typ_subst_atomic; val Ds0 = map permute_Ds Ds0; (* terms for sets of the set_transfer thms *) val rel_sets_Q = @{map 3} (fn aT => fn bT => fn Q => mk_rel 1 [aT] [bT] @{term rel_set} $ Q) (#alphas Tss) (#betas Tss) var_Qs; (* rewrite rules for pcrel and BNF's relator: "pcrel Q = rel_F OO crel" *) fun mk_pcr_rel_F_eq Ts Us pcrel vs crel = let val thm = HOLogic.mk_Trueprop (HOLogic.mk_eq (pcrel, mk_relcompp (list_comb (mk_rel_of_bnf (#deads Tss) (Ts Tss) (Us Tss) bnf_F, vs)) crel)); fun tac ctxt = unfold_thms_tac ctxt (pcrel_def :: defs @ @{thm id_bnf_apply} :: Transfer.get_relator_eq ctxt) THEN (HEADGOAL (rtac ctxt refl)) in prove lthy vs thm tac |> mk_abs_def end; val pcr_rel_F_eqs = [mk_pcr_rel_F_eq #alphas #betas pcrel_Qs var_Qs crel_b, mk_pcr_rel_F_eq #gammas #deltas pcrel_Rs var_Rs crel_d]; (* create transfer-relations from term('s type) *) fun mk_transfer_rel' i tm = let fun go T' (n, q) = case T' of Type ("fun", [T as Type ("fun", _), U]) => go U (n+1, q) |>> mk_rel_fun (fst (go T (n, q))) | Type ("fun", [T, U]) => go T (n, q) |-> (fn x => fn st => go U st |>> mk_rel_fun x) | Type (@{type_name bool}, _) => (HOLogic.eq_const HOLogic.boolT, (n, q)) | Type (@{type_name set}, _) => (nth rel_sets_Q n, (n, q)) | Type _ => (if q then pcrel_Qs else pcrel_Rs, (n, not q)) | TFree _ => (nth (if q then var_Qs else var_Rs) n, (n, not q)) | _ => raise Match in go (fastype_of tm) (i, true) |> fst end; (* prove transfer rules *) fun prove_transfer_thm' i vars new_tm const = let val tm = HOLogic.mk_Trueprop (mk_transfer_rel' i new_tm $ #raw const $ new_tm); val tac = (fn ctxt => unfold_thms_tac ctxt (pcr_rel_F_eqs @ [crel_def]) THEN #tac const {Rs=var_Rs,Qs=var_Qs} ctxt); in prove lthy vars tm tac end; val prove_transfer_thm = prove_transfer_thm' 0; (* map transfer: "((Q ===> R) ===> pcr_G Q ===> pcr_G R) map_F' map_G" *) val map_G = mk_map_of_bnf Ds0 (#betas Tss) (#deltas Tss) bnf_G; val map_transfer = prove_transfer_thm (var_Qs @ var_Rs) map_G (#map consts); (* pred_transfer: "((Q ===> (=)) ===> pcr_G Q ===> (=)) pred_F' pred_G" *) val pred_G = mk_pred_of_bnf Ds0 (#betas Tss) bnf_G; val pred_transfer = #pred consts |> Option.map (fn p => ("pred", [prove_transfer_thm (var_Qs @ var_Rs) pred_G p])); (* rel_transfer: "((Q ===> R ===> (=)) ===> pcr_G Q ===> pcr_G R ===> (=)) rel_F' rel_G" *) val rel_G = mk_rel_of_bnf Ds0 (#betas Tss) (#deltas Tss) bnf_G; val rel_transfer = prove_transfer_thm (var_Qs @ var_Rs) rel_G (#rel consts); (* set_transfer: "(pcr_G Q ===> rel_set Q) set_F' set_G" *) val sets_G = mk_sets_of_bnf (replicate live Ds0) (replicate live (#betas Tss)) bnf_G; fun mk_set_transfer i set_G raw tac = prove_transfer_thm' i var_Qs set_G {raw=raw, tac=tac}; val sets_transfer = @{map 4} mk_set_transfer (0 upto (live-1)) sets_G (#raws (#sets consts)) (#tacs (#sets consts)); (* export transfer theorems *) val transform = Morphism.thm (Morphism.thm_morphism "BNF" (unfold_thms lthy defs)) |> map; val b = Binding.qualified_name name; val qualify = let val qs = Binding.path_of b; in fold_rev (fn (s, mand) => Binding.qualify mand s) qs end; fun mk_binding n = Binding.name (n ^ "_transfer_raw") |> Binding.qualify true (Binding.name_of b) |> qualify; val notes = [("map", [map_transfer]), ("rel", [rel_transfer])] @ the_list pred_transfer @ [("set", sets_transfer)] |> map (fn (n, thms) => ((mk_binding n, []), [(transform thms, @{attributes [transfer_rule]})])); in lthy |> Local_Theory.notes notes |> snd end; (* transfer theorems for map, pred (in case of a typedef), rel and sets *) fun mk_transfer_thms quiet bnf_F bnf_G name consts thm Tss defs lthy = let fun mk_crel_def quot_thm = (case thm of Quotient equiv => @{thm Quotient_crel_quotient} OF [quot_thm, equiv] - | Typedef _ => hd ([quot_thm] RL @{thms Quotient_crel_typedef Quotient_crel_typecopy})); + | Typedef => hd ([quot_thm] RL @{thms Quotient_crel_typedef Quotient_crel_typecopy})); fun no_quotient _ = [Pretty.para ("No quotient theorem has been registered for " ^ Binding.name_of (name_of_bnf bnf_G) ^ "."), Pretty.para "Use setup_lifting to register a quotient or type definition theorem."]; fun wrong_quotient T lthy = [Pretty.para ("A wrong quotient theorem has been registered for " ^ Binding.name_of (name_of_bnf bnf_G) ^ "."), Pretty.para ("Expected raw type " ^ Pretty.string_of (Syntax.pretty_typ lthy (T_of_bnf bnf_F)) ^ " but the quotient theorem has raw type " ^ Pretty.string_of (Syntax.pretty_typ lthy T) ^ "."), Pretty.para "Use setup_lifting to register a different quotient or type definition theorem."]; fun pcr_why _ = [Pretty.para ("The pcr_" ^ Binding.name_of (name_of_bnf bnf_G) ^ " relator has not been defined.")]; fun warn_transfer why lthy = (Pretty.para "The transfer theorems can't be generated:" :: why lthy) |> Pretty.chunks |> Pretty.string_of |> warning |> K lthy; fun maybe_warn_transfer why = not quiet ? warn_transfer why; in case Lifting_Info.lookup_quotients lthy name of SOME {pcr_info, quot_thm} => (let val crel_def = mk_crel_def quot_thm; val rty = Lifting_Util.quot_thm_rty_qty quot_thm |> fst; val thy = Proof_Context.theory_of lthy; in if Sign.typ_instance thy (rty, T_of_bnf bnf_F) then (case pcr_info of SOME {pcrel_def, ...} => mk_transfer_thms' bnf_F bnf_G name consts Tss crel_def pcrel_def defs lthy | _ => maybe_warn_transfer pcr_why lthy) else maybe_warn_transfer (wrong_quotient rty) lthy end) | _ => maybe_warn_transfer no_quotient lthy end; (** typedef_bnf **) fun mk_typedef_transfer_tacs bnf_F bnf_G thms old_defs map_raw rel_raw pred_raw sets_raw = let val live = live_of_bnf bnf_G; val Abs_G_inverse = @{thm type_definition.Abs_inverse} OF [#typedef thms]; val Rep_G = @{thm type_definition.Rep} OF [#typedef thms]; fun common_tac addefs tac = (fn _ => fn ctxt => HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt addefs), REPEAT_DETERM o rtac ctxt rel_funI, SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply}), REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE exE conjE}, hyp_subst_tac ctxt]) THEN tac ctxt) fun map_tac ctxt = (HEADGOAL o EVERY') [rtac ctxt @{thm relcomppI}, etac ctxt (mk_rel_funDN (live+1) (map_transfer_of_bnf bnf_F)), REPEAT_DETERM_N live o assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [Abs_G_inverse OF [#map_closed thms] OF [Rep_G]]), REPEAT_DETERM o rtac ctxt refl]; val map_tac = common_tac [#map old_defs] map_tac; fun rel_tac ctxt = HEADGOAL (etac ctxt (mk_rel_funDN (live+2) (rel_transfer_of_bnf bnf_F)) THEN' REPEAT_DETERM_N (live+1) o assume_tac ctxt); val rel_tac = common_tac (#rel old_defs :: @{thms vimage2p_def}) rel_tac; fun pred_tac ctxt = HEADGOAL (etac ctxt (mk_rel_funDN (live+1) (pred_transfer_of_bnf bnf_F)) THEN' REPEAT_DETERM_N live o (assume_tac ctxt)); val pred_tac = common_tac [#pred old_defs] pred_tac; fun set_tac set_transfer_thm ctxt = HEADGOAL (etac ctxt (rel_funD OF [set_transfer_thm])); fun mk_set_tac set_def set_transfer = common_tac [set_def] (set_tac set_transfer); val set_tacs = map2 mk_set_tac (#sets old_defs) (set_transfer_of_bnf bnf_F); in {map={raw=map_raw,tac=map_tac},rel={raw=rel_raw,tac=rel_tac}, sets={raws=sets_raw,tacs=set_tacs},pred=SOME{raw=pred_raw,tac=pred_tac}} end; fun typedef_bnf thm wits specs map_b rel_b pred_b opts lthy = let val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; (* extract Rep Abs F RepT AbsT *) val (Rep_G, Abs_G, F) = strip3 thm; val typ_Abs_G = dest_funT (fastype_of Abs_G); val RepT = fst typ_Abs_G; (* F *) val AbsT = snd typ_Abs_G; (* G *) val AbsT_name = fst (dest_Type AbsT); val tvs = AbsT |> dest_Type |> snd |> map (fst o dest_TVar); val alpha0s = map (TFree o snd) specs; val _ = length tvs = length alpha0s orelse error ("Expected " ^ string_of_int (length tvs) ^ " type argument(s) to " ^ quote AbsT_name); (* instantiate the new type variables newtvs to oldtvs *) val subst = subst_TVars (tvs ~~ alpha0s); val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); val Rep_G = subst Rep_G; val Abs_G = subst Abs_G; val F = subst F; val RepT = typ_subst RepT; val AbsT = typ_subst AbsT; fun flatten_tyargs Ass = map dest_TFree alpha0s |> filter (fn T => exists (fn Ts => member op= Ts T) Ass); val Ds0 = filter (is_none o fst) specs |> map snd; (* get the bnf for RepT *) val ((bnf_F, (deads, alphas)), ((_, unfolds), lthy)) = bnf_of_typ true Dont_Inline (Binding.qualify true AbsT_name) flatten_tyargs [] Ds0 RepT ((empty_comp_cache, empty_unfolds), lthy); val set_bs = map (fn T => find_index (fn U => T = U) alpha0s) alphas |> map (the_default Binding.empty o fst o nth specs); val _ = (case alphas of [] => error "No live variables" | _ => ()); val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds @ #pred_unfolds unfolds; (* number of live variables *) val live = length alphas; (* state the three required properties *) val sorts = map Type.sort_of_atyp alphas; val names_lthy = fold Variable.declare_typ (alphas @ deads) lthy; val (((alphas', betas), betas'), names_lthy) = names_lthy |> mk_TFrees' sorts ||>> mk_TFrees' sorts ||>> mk_TFrees' sorts; val map_F = mk_map_of_bnf deads alphas betas bnf_F; val (typ_fs, typ_aF) = fastype_of map_F |> strip_typeN live ||> domain_type; val typ_pairs = map HOLogic.mk_prodT (alphas ~~ alphas'); val typ_subst_pair = typ_subst_atomic (alphas ~~ typ_pairs); val typ_pair = typ_subst_pair RepT; val subst_b = subst_atomic_types (alphas ~~ betas); val subst_a' = subst_atomic_types (alphas ~~ alphas'); val subst_pair = subst_atomic_types (alphas ~~ typ_pairs); val aF_set = F; val aF_set' = subst_a' F; val pairF_set = subst_pair F; val bF_set = subst_b F; val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf_F; val map_F_snd = mk_map_of_bnf deads typ_pairs alphas' bnf_F val sets_F_pairs = mk_sets_of_bnf (replicate live deads) (replicate live typ_pairs) bnf_F val wits_F = mk_wits_of_bnf (replicate (nwits_of_bnf bnf_F) deads) (replicate (nwits_of_bnf bnf_F) alphas) bnf_F; (* val map_closed_F = @{term "\f x. x \ F \ map_F f x \ F"}; *) val (var_fs, names_lthy) = mk_Frees "f" typ_fs names_lthy; val (var_x, names_lthy) = mk_Frees "x" [typ_aF] names_lthy |>> the_single; val mem_x = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_x, aF_set)); val map_f = list_comb (map_F, var_fs); val mem_map = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_f $ var_x, bF_set)); val imp_map = Logic.mk_implies (mem_x, mem_map); val map_closed_F = fold_rev Logic.all var_fs (Logic.all var_x imp_map); (* val zip_closed_F = @{term "\z. \map_F fst z \ F; map_F snd z \ F\ \ \z' \ F. set_F z' \ set_F z \ map_F fst z' = map_F fst z \ map_F snd z' = map_F snd z"}; *) val (var_z, names_lthy) = mk_Free "z" typ_pair names_lthy; val (var_z', names_lthy) = mk_Free "z'" typ_pair names_lthy; val (pairs, names_lthy) = mk_Frees "tmp" typ_pairs names_lthy; fun mk_map mfs f z = Term.list_comb (mfs, map (fst o Term.strip_comb o f) pairs) $ z; fun mk_set var = map (fn t => t $ var) sets_F_pairs; val (map_fst', map_fst) = apply2 (mk_map map_F_fst HOLogic.mk_fst) (var_z', var_z); val (map_snd', map_snd) = apply2 (mk_map map_F_snd HOLogic.mk_snd) (var_z', var_z); val mem_map_fst = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_fst, aF_set)); val mem_map_snd = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_snd, aF_set')); val ex_conj = foldr1 HOLogic.mk_conj (map2 mk_leq (mk_set var_z') (mk_set var_z) @ [HOLogic.mk_eq (map_fst', map_fst), HOLogic.mk_eq (map_snd', map_snd)]); val zip_concl = HOLogic.mk_Trueprop (mk_Bex pairF_set (absfree (dest_Free var_z') ex_conj)); val zip_closed_F = Logic.all var_z (Logic.list_implies ([mem_map_fst, mem_map_snd], zip_concl)); (* val wit_closed_F = @{term "wit_F a \ F"}; *) val (var_as, names_lthy) = mk_Frees "a" alphas names_lthy; val (var_bs, _) = mk_Frees "a" alphas names_lthy; val sets = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val (Iwits, wit_goals) = prepare_wits false RepT wits opts alphas wits_F var_as var_bs sets lthy; val wit_closed_Fs = Iwits |> map (fn (I, wit_F) => let val vars = map (nth var_as) I; val wit_a = list_comb (wit_F, vars); in fold_rev Logic.all vars (HOLogic.mk_Trueprop (HOLogic.mk_mem (wit_a, aF_set))) end); val goals = [map_closed_F, zip_closed_F] @ wit_closed_Fs @ (case wits of NONE => [] | _ => wit_goals); fun after_qed ([map_closed_thm] :: [zip_closed_thm] :: wit_thmss) lthy = let val (wit_closed_thms, wit_thms) = (case wits of NONE => (map the_single wit_thmss, wit_thms_of_bnf bnf_F) | _ => chop (length wit_closed_Fs) (map the_single wit_thmss)) (* construct map set bd rel wit *) (* val map_G = @{term "\f. Abs_G o map_F f o Rep_G"}; *) val Abs_Gb = subst_b Abs_G; val map_G = fold_rev lambda var_fs (HOLogic.mk_comp (HOLogic.mk_comp (Abs_Gb, map_f), Rep_G)); val map_raw = fold_rev lambda var_fs map_f; (* val sets_G = [@{term "set_F o Rep_G"}]; *) val sets_F = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val sets_G = map (fn set_F => HOLogic.mk_comp (set_F, Rep_G)) sets_F; (* val bd_G = @{term "bd_F"}; *) val bd_F = mk_bd_of_bnf deads alphas bnf_F; val bd_G = bd_F; (* val rel_G = @{term "\R. BNF_Def.vimage2p Rep_G Rep_G (rel_F R)"}; *) val rel_F = mk_rel_of_bnf deads alphas betas bnf_F; val (typ_Rs, _) = strip_typeN live (fastype_of rel_F); val (var_Rs, names_lthy) = mk_Frees "R" typ_Rs lthy; val Rep_Gb = subst_b Rep_G; val rel_G = fold_rev absfree (map dest_Free var_Rs) (mk_vimage2p Rep_G Rep_Gb $ list_comb (rel_F, var_Rs)); val rel_raw = fold_rev absfree (map dest_Free var_Rs) (list_comb (rel_F, var_Rs)); (* val pred_G = @{term "\P. pred_F P o Rep_G"}; *) val pred_F = mk_pred_of_bnf deads alphas bnf_F; val (typ_Ps, _) = strip_typeN live (fastype_of pred_F); val (var_Ps, names_lthy) = mk_Frees "P" typ_Ps names_lthy; val pred_G = fold_rev absfree (map dest_Free var_Ps) (HOLogic.mk_comp (list_comb (pred_F, var_Ps), Rep_G)); val pred_raw = fold_rev absfree (map dest_Free var_Ps) (list_comb (pred_F, var_Ps)); (* val wits_G = [@{term "Abs_G o wit_F"}]; *) val (var_as, _) = mk_Frees "a" alphas names_lthy; val wits_G = map (fn (I, wit_F) => let val vs = map (nth var_as) I; in fold_rev absfree (map dest_Free vs) (Abs_G $ (list_comb (wit_F, vs))) end) Iwits; (* tactics *) val Rep_thm = thm RS @{thm type_definition.Rep}; val Abs_inverse_thm = thm RS @{thm type_definition.Abs_inverse}; val Abs_inject_thm = thm RS @{thm type_definition.Abs_inject}; val Rep_cases_thm = thm RS @{thm type_definition.Rep_cases}; val Rep_inverse_thm = thm RS @{thm type_definition.Rep_inverse}; fun map_id0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, SELECT_GOAL (unfold_thms_tac ctxt [map_id0_of_bnf bnf_F, id_apply, o_apply, Rep_inverse_thm]), rtac ctxt refl]); fun map_comp0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, SELECT_GOAL (unfold_thms_tac ctxt [map_comp0_of_bnf bnf_F, o_apply, Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), rtac ctxt refl]); fun map_cong0_tac ctxt = HEADGOAL (EVERY' ([SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt (([Rep_thm RS map_closed_thm, Rep_thm RS map_closed_thm] MRS Abs_inject_thm) RS iffD2), rtac ctxt (map_cong0_of_bnf bnf_F)] @ replicate live (Goal.assume_rule_tac ctxt))); val set_map0s_tac = map (fn set_map => fn ctxt => HEADGOAL (EVERY' [rtac ctxt ext, SELECT_GOAL (unfold_thms_tac ctxt [set_map, o_apply, Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), rtac ctxt refl])) (set_map_of_bnf bnf_F); fun card_order_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_card_order_of_bnf bnf_F)); fun cinfinite_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_cinfinite_of_bnf bnf_F)); val set_bds_tac = map (fn set_bd => fn ctxt => HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt set_bd])) (set_bd_of_bnf bnf_F); fun le_rel_OO_tac ctxt = HEADGOAL (EVERY' [rtac ctxt @{thm vimage2p_relcompp_mono}, rtac ctxt ((rel_OO_of_bnf bnf_F RS sym) RS @{thm ord_eq_le_trans}), rtac ctxt @{thm order_refl}]); fun rel_OO_Grp_tac ctxt = HEADGOAL (EVERY' ([SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt ext))), SELECT_GOAL (unfold_thms_tac ctxt [@{thm OO_Grp_alt}, mem_Collect_eq, o_apply, @{thm vimage2p_def}, in_rel_of_bnf bnf_F, Bex_def, mem_Collect_eq]), rtac ctxt iffI, SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [exE, conjE]))), forward_tac ctxt [zip_closed_thm OF (replicate 2 (Rep_thm RSN (2, @{thm ssubst_mem})))], assume_tac ctxt, SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [bexE, conjE]))), etac ctxt Rep_cases_thm, hyp_subst_tac ctxt, SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt exI))), rtac ctxt conjI] @ replicate live (EVERY' [TRY o rtac ctxt conjI, etac ctxt @{thm order_trans}, assume_tac ctxt]) @ [SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt conjI))), REPEAT_DETERM_N 2 o EVERY' [rtac ctxt (trans OF [iffD2 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]], Rep_inverse_thm]), etac ctxt trans, assume_tac ctxt], SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [exE, conjE]))), rtac ctxt exI, rtac ctxt conjI] @ replicate (live-1) (rtac ctxt conjI THEN' assume_tac ctxt) @ [assume_tac ctxt, rtac ctxt conjI, REPEAT_DETERM_N 2 o EVERY' [rtac ctxt (iffD1 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]]), etac ctxt (Rep_inverse_thm RS sym RSN (2, trans))]])); fun pred_set_tac ctxt = HEADGOAL (EVERY' [rtac ctxt (pred_set_of_bnf bnf_F RS @{thm arg_cong[of _ _ "\f. f \ _"]} RS trans), SELECT_GOAL (unfold_thms_tac ctxt (@{thms Ball_comp_iff conj_comp_iff})), rtac ctxt refl]); fun wit_tac ctxt = HEADGOAL (EVERY' (map (fn thm => (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt (o_apply :: (wit_closed_thms RL [Abs_inverse_thm]))), dtac ctxt thm, assume_tac ctxt])) wit_thms)); val tactics = [map_id0_tac, map_comp0_tac, map_cong0_tac] @ set_map0s_tac @ [card_order_bd_tac, cinfinite_bd_tac] @ set_bds_tac @ [le_rel_OO_tac, rel_OO_Grp_tac, pred_set_tac]; val (bnf_G, lthy) = bnf_def Dont_Inline (user_policy Note_Some) true I tactics wit_tac NONE map_b rel_b pred_b set_bs (((((((Binding.empty, AbsT), map_G), sets_G), bd_G), wits_G), SOME rel_G), SOME pred_G) lthy; val old_defs = {sets = set_defs_of_bnf bnf_G, map = map_def_of_bnf bnf_G, rel = rel_def_of_bnf bnf_G, pred = pred_def_of_bnf bnf_G}; val unfold_morphism = Morphism.thm_morphism "BNF" (unfold_thms lthy defs); val (bnf_G, lthy) = morph_bnf_defs unfold_morphism bnf_G |> (fn bnf => note_bnf_defs bnf lthy); - val setup_lifting_thm = Typedef thm; val quiet = exists (fn No_Warn_Transfer => true | _ => false) opts; val transfer_consts = mk_typedef_transfer_tacs bnf_F bnf_G {map_closed=map_closed_thm,typedef=thm} old_defs map_raw rel_raw pred_raw sets_F; in lthy |> BNF_Def.register_bnf plugins AbsT_name bnf_G |> - mk_transfer_thms quiet bnf_F bnf_G AbsT_name transfer_consts setup_lifting_thm + mk_transfer_thms quiet bnf_F bnf_G AbsT_name transfer_consts Typedef {abs=typ_subst_atomic (alphas ~~ alphas') AbsT, rep=RepT, Ds0=map TFree Ds0, deads = deads, alphas=alphas, betas=alphas', gammas=betas, deltas=betas'} defs end | after_qed _ _ = raise Match; in (goals, after_qed, defs, lthy) end; (** quotient_bnf **) fun mk_quotient_transfer_tacs bnf_F Tss live qthms thms set_F'_thmss old_defs inst_REL_pos_distrI map_raw rel_raw sets_raw = let fun common_tac ctxt addefs = unfold_thms_tac ctxt (#REL qthms :: addefs) THEN (REPEAT_DETERM o HEADGOAL) (rtac ctxt rel_funI); (* quotient.map_transfer tactic *) val map_F_transfer = map_transfer_of_bnf bnf_F |> mk_rel_funDN (live+1); fun map_transfer_q _ ctxt = common_tac ctxt (#map old_defs :: @{thms o_def}) THEN (HEADGOAL o EVERY') [REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE}, rtac ctxt @{thm relcomppI[rotated]}, hyp_subst_tac ctxt THEN' EVERY' (map (rtac ctxt) [#rel_abs qthms, #map_F_rsp thms, (#rep_abs_rsp qthms), (#reflp qthms)]), hyp_subst_tac ctxt, rtac ctxt map_F_transfer, REPEAT_DETERM_N (live+1) o assume_tac ctxt]; (* quotient.rel_transfer tactic *) val rel_F_maps = rel_map_of_bnf bnf_F; val rel_F_map_iffD2s = map (fn thm => thm RS @{thm iffD2}) rel_F_maps; fun inst_REL_pos_distrI_order_refls vs aTs bTs ctxt = inst_REL_pos_distrI live vs aTs bTs ctxt OF (replicate (live+1) asm_rl @ replicate live @{thm order_refl}); fun rel_transfer_q {Qs, Rs} ctxt = EVERY [common_tac ctxt [#rel old_defs, @{thm vimage2p_def}], HEADGOAL (rtac ctxt iffI), (REPEAT_DETERM o ALLGOALS) (eresolve_tac ctxt @{thms exE conjE relcomppE} ORELSE' hyp_subst_tac ctxt), (HEADGOAL o EVERY') [REPEAT_DETERM o dtac ctxt @{thm rel_fun_rel_OO1}, rtac ctxt (inst_REL_pos_distrI 0 (map mk_conversep Qs) (#betas Tss) (#alphas Tss) ctxt), rtac ctxt @{thm relcomppI}, rtac ctxt (#symp qthms), rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms), rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI}, rtac ctxt (rel_flip_of_bnf bnf_F RS @{thm iffD1}), rtac ctxt (nth rel_F_map_iffD2s 0), rtac ctxt (nth rel_F_map_iffD2s 1), etac ctxt (#rel_monoD_rotated thms)], (REPEAT_DETERM_N live o HEADGOAL o EVERY') [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm Basic_BNFs.rel_sum_simps(4)[THEN iffD2]}, etac ctxt @{thm conversepI}], (HEADGOAL o EVERY') [rtac ctxt (inst_REL_pos_distrI_order_refls Rs (#gammas Tss) (#deltas Tss) ctxt), (SELECT_GOAL o REPEAT_DETERM o HEADGOAL) (etac ctxt @{thm relcomppI}), rtac ctxt @{thm relcomppI[rotated]}, rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms OF [#reflp qthms]), SELECT_GOAL (unfold_thms_tac ctxt (@{thms rel_sum_simps} @ rel_F_maps)), assume_tac ctxt], (REPEAT_DETERM_N (2*live) o HEADGOAL) (rtac ctxt @{thm rel_sum_eq2_nonempty} ORELSE' rtac ctxt @{thm rel_sum_eq3_nonempty}), (REPEAT_DETERM_N live) (unfold_thms_tac ctxt @{thms sum.rel_compp[symmetric] eq_OO} THEN HEADGOAL (etac ctxt @{thm sum.rel_mono[OF order_refl]})), (HEADGOAL o EVERY') [(SELECT_GOAL o REPEAT_DETERM o HEADGOAL) (dtac ctxt @{thm rel_fun_rel_OO2}), rtac ctxt (inst_REL_pos_distrI 0 Qs (#alphas Tss) (#betas Tss) ctxt), rtac ctxt @{thm relcomppI}, rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI}, rtac ctxt (nth rel_F_map_iffD2s 0), rtac ctxt (nth rel_F_map_iffD2s 1), etac ctxt (#rel_monoD_rotated thms)], (REPEAT_DETERM_N live o HEADGOAL o EVERY') [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm rel_sum.intros(2)}, assume_tac ctxt], (HEADGOAL o EVERY') [rtac ctxt (inst_REL_pos_distrI_order_refls (map mk_conversep Rs) (#deltas Tss) (#gammas Tss) ctxt), rtac ctxt @{thm relcomppI}, etac ctxt (rotate_prems 1 (#transp qthms)), rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms OF [#reflp qthms]), etac ctxt @{thm relcomppI}, rtac ctxt @{thm relcomppI}, etac ctxt (#transp qthms), rtac ctxt (#symp qthms), rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms), rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI[rotated]}, rtac ctxt (#reflp qthms), rtac ctxt (rel_flip_of_bnf bnf_F RS @{thm iffD1}), rtac ctxt (nth rel_F_map_iffD2s 0), rtac ctxt (nth rel_F_map_iffD2s 1), etac ctxt (#rel_monoD_rotated thms)], (REPEAT_DETERM_N live o HEADGOAL o EVERY') [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm rel_sum.intros(2)}, etac ctxt @{thm conversepI}], (REPEAT_DETERM_N (2*live) o HEADGOAL) (rtac ctxt @{thm rel_sum_eq2_nonempty} ORELSE' rtac ctxt @{thm rel_sum_eq3_nonempty}), (REPEAT_DETERM_N live o EVERY) [unfold_thms_tac ctxt @{thms sum.rel_compp[symmetric] eq_OO}, HEADGOAL (etac ctxt @{thm sum.rel_mono[OF order_refl]})]]; (* quotient.set_transfer tactics *) fun set_transfer_q set_G_def set_F'_thms _ ctxt = let val set_F'_rsp = mk_rel_funDN 1 (#set_F'_respect set_F'_thms) in common_tac ctxt (set_G_def :: @{thms o_def}) THEN (HEADGOAL o EVERY') [etac ctxt @{thm relcomppE}, hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [set_F'_rsp OF [#rep_abs qthms] OF [#reflp qthms], @{thm rel_set_def}]), dtac ctxt (#rel_F_rel_F' thms), rtac ctxt conjI] THEN (REPEAT_DETERM_N 2 o HEADGOAL o EVERY') [SELECT_GOAL (unfold_thms_tac ctxt [#rel_F'_set thms]), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], REPEAT_DETERM o dtac ctxt (mk_sym set_F'_rsp), SELECT_GOAL (unfold_thms_tac ctxt [#set_map_F' set_F'_thms]), rtac ctxt ballI, dtac ctxt @{thm equalityD1[THEN subsetD]}, assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms image_iff}), etac ctxt bexE, dtac ctxt set_mp, assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms mem_Collect_eq case_prod_beta}), rtac ctxt bexI, hyp_subst_tac ctxt, assume_tac ctxt, etac ctxt @{thm hypsubst}, etac ctxt @{thm imageI}, assume_tac ctxt] end; in {map={raw=map_raw, tac=map_transfer_q}, rel={raw=rel_raw, tac=rel_transfer_q}, sets={raws=sets_raw,tacs=map2 set_transfer_q (#sets old_defs) set_F'_thmss}, pred=NONE} end; -fun quotient_bnf (equiv_thm, quot_thm) _ wits specs map_b rel_b pred_b opts lthy = +fun quotient_bnf (equiv_thm, quot_thm) wits specs map_b rel_b pred_b opts lthy = let (* extract rep_G and abs_G *) val (equiv_rel, abs_G, rep_G) = strip3 quot_thm; val (repT, absT) = dest_funT (fastype_of abs_G); (* ("?'a F", "?'a G") *) val absT_name = fst (dest_Type absT); val tvs = map (fst o dest_TVar) (snd (dest_Type absT)); val _ = length tvs = length specs orelse error ("Expected " ^ string_of_int (length tvs) ^ " type argument" ^ (if (length tvs) = 1 then "" else "s") ^ " to " ^ quote absT_name); (* instantiate TVars *) val alpha0s = map (TFree o snd) specs; val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); val (repT, absT) = apply2 typ_subst (repT, absT); (* get the bnf for RepT *) val Ds0 = filter (is_none o fst) specs |> map snd; fun flatten_tyargs Ass = map dest_TFree alpha0s |> filter (fn T => exists (fn Ts => member op= Ts T) Ass); val ((bnf_F, (deads, alphas)), ((_, unfolds), lthy)) = bnf_of_typ true Dont_Inline (Binding.qualify true absT_name) flatten_tyargs [] Ds0 repT ((empty_comp_cache, empty_unfolds), lthy); val live = length alphas; val _ = (if live = 0 then error "No live variables" else ()); val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds; val set_bs = map (fn T => find_index (fn U => T = U) alpha0s) alphas |> map (the_default Binding.empty o fst o nth specs); (* create and instantiate all the needed type variables *) val subst = subst_TVars (tvs ~~ alpha0s); val (abs_G, rep_G) = apply2 subst (abs_G, rep_G); val sorts = map Type.sort_of_atyp alphas; val (((betas, gammas), deltas), names_lthy) = fold Variable.declare_typ (alphas @ deads) lthy |> mk_TFrees' sorts ||>> mk_TFrees' sorts ||>> mk_TFrees' sorts; fun subst_Ts tm Ts = subst_atomic_types (alphas ~~ Ts) tm; val subst_b = subst_atomic_types (alphas ~~ betas); val subst_Maybe = subst_atomic_types o map (swap o `mk_MaybeT); val equiv_rel_a = subst equiv_rel; val map_F = mk_map_of_bnf deads alphas betas bnf_F; val rel_F_ab = mk_rel_of_bnf deads alphas betas bnf_F; val rel_F_bc = mk_rel_of_bnf deads betas gammas bnf_F; val rel_F_ac = mk_rel_of_bnf deads alphas gammas bnf_F; val rel_F_option = mk_rel_of_bnf deads (map mk_MaybeT alphas) (map mk_MaybeT betas) bnf_F; val sets_F = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val wits_F = mk_wits_of_bnf (replicate (nwits_of_bnf bnf_F) deads) (replicate (nwits_of_bnf bnf_F) alphas) bnf_F; val (typ_fs, (typ_aF, typ_bF)) = strip_typeN live (fastype_of map_F) ||> dest_funT; val typ_MaybeF = typ_subst_atomic (alphas ~~ map mk_MaybeT alphas) typ_aF; val typ_a_sets = map HOLogic.mk_setT alphas; val typ_pairs = map HOLogic.mk_prodT (alphas ~~ betas); val typ_fs' = map (typ_subst_atomic (map (swap o `mk_MaybeT) betas)) typ_fs; (* create all the needed variables *) val ((((((((((((((((((((((var_Ps, var_Qs), var_Rs), var_x), var_x'), var_y), var_y'), var_mx), var_As), var_As'), var_Ss), var_Bs), var_as), var_as'), var_bs), var_bs'), var_R), var_fs), var_fs'), var_gs), var_gs'), var_z), var_ts) = names_lthy |> mk_Frees "Ps" (map2 mk_relT alphas betas) ||>> mk_Frees "Qs" (map2 mk_relT betas gammas) ||>> mk_Frees "Rs" (map2 mk_relT alphas gammas) ||>> mk_Free "x" typ_aF ||>> mk_Free "x'" typ_aF ||>> mk_Free "y" typ_bF ||>> mk_Free "y'" (typ_subst_atomic (alphas ~~ gammas) typ_aF) ||>> mk_Free "mx" typ_MaybeF ||>> mk_Frees "As" typ_a_sets ||>> mk_Frees "As'" typ_a_sets ||>> mk_Frees "Ss" (map HOLogic.mk_setT typ_a_sets) ||>> mk_Frees "Bs" (map HOLogic.mk_setT betas) ||>> mk_Frees "as" alphas ||>> mk_Frees "as'" alphas ||>> mk_Frees "bs" betas ||>> mk_Frees "bs'" betas ||>> mk_Free "R" (typ_aF --> typ_bF --> HOLogic.boolT) ||>> mk_Frees "fs" typ_fs ||>> mk_Frees "fs'" typ_fs' ||>> mk_Frees "gs" typ_fs ||>> mk_Frees "gs'" typ_fs' ||>> mk_Free "z" (typ_subst_atomic (alphas ~~ typ_pairs) typ_aF) ||>> mk_Frees "ts" typ_pairs |> fst; (* create local definitions `b = tm` with n arguments *) fun suffix tm s = (dest_Const tm |> fst |> Long_Name.base_name) ^ s; fun define lthy b n tm = let val b = Binding.qualify true absT_name (Binding.qualified_name b); val ((tm, (_, def)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define_internal (((Binding.concealed b, NoSyn), (Binding.empty_atts, tm))) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val tm = Term.subst_atomic_types (map (`(Morphism.typ phi)) (alphas @ betas @ gammas @ map TFree Ds0)) (Morphism.term phi tm); val def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def)); in ({def=def, tm=tm}, lthy) end; (* internally use REL, not the user-provided definition *) val (REL, lthy) = define lthy "REL" 0 equiv_rel_a; val REL_def = sym RS eq_reflection OF [#def REL]; fun REL_rewr_all ctxt thm = Conv.fconv_rule (Conv.top_conv (fn _ => Conv.try_conv (Conv.rewr_conv REL_def)) ctxt) thm; val equiv_rel_a' = equiv_rel_a; val equiv_rel_a = #tm REL; val (equiv_rel_b, equiv_rel_c) = apply2 (subst_Ts equiv_rel_a) (betas, gammas); (* rel_pos_distr: @{term "\A B. A \\ B \ bot \ REL \\ rel_F A \\ REL \\ rel_F B \\ REL \ REL \\ rel_F (A \\ B) \\ REL"} *) fun compp_not_bot comp aT cT = let val T = mk_relT aT cT; val mk_eq = HOLogic.eq_const T; in HOLogic.mk_not (mk_eq $ comp $ bot_const T) end; val ab_comps = map2 mk_relcompp var_Ps var_Qs; val ne_comps = (@{map 3} compp_not_bot ab_comps alphas gammas); val ab_prem = foldr1 HOLogic.mk_conj ne_comps; val REL_pos_distrI_tm = let val le_relcomps = map2 mk_leq ab_comps var_Rs; val assm = mk_OO [equiv_rel_a, list_comb (rel_F_ab, var_Ps), equiv_rel_b, list_comb (rel_F_bc, var_Qs)] equiv_rel_c; val concl = mk_OO [equiv_rel_a, list_comb (rel_F_ac, var_Rs)] equiv_rel_c; in mk_Trueprop_implies ([assm $ var_x $ var_y'] @ ne_comps @ le_relcomps, concl $ var_x $ var_y') end; val ab_concl = mk_leq (mk_OO [list_comb (rel_F_ab, var_Ps), equiv_rel_b] (list_comb (rel_F_bc, var_Qs))) (mk_OO [equiv_rel_a, list_comb (rel_F_ac, ab_comps)] (equiv_rel_c)); val ab_imp = Logic.mk_implies (apply2 HOLogic.mk_Trueprop (ab_prem, ab_concl)); val rel_pos_distr = fold_rev Logic.all (var_Ps @ var_Qs) ab_imp; (* {(x, y) . REL x y} *) fun mk_rel_pairs rel = mk_case_prod (var_x, var_x') (rel $ var_x $ var_x') val rel_pairs = mk_rel_pairs equiv_rel_a; (* rel_Inter: \S. \ S \ {}; \S \ {} \ \ (\A\S. {(x, y). REL x y} `` {x. set_F x \ A}) \ {(x, y). REL x y} `` {x. set_F x \ \S} *) fun rel_Inter_from_set_F (var_A, var_S) set_F = let val typ_aset = fastype_of var_A; (* \S *) val inter_S = Inf_const typ_aset $ var_S; (* [S \ {}, \S \ {}] *) fun not_empty x = let val ty = fastype_of x in HOLogic.mk_not (HOLogic.mk_eq (x, bot_const ty)) end; val prems = map (HOLogic.mk_Trueprop o not_empty) [var_S, inter_S]; (* {x. set_F x \ A} *) val setF_sub_A = mk_in [var_A] [set_F] typ_aF; (* {x. set_F x \ \S} *) val setF_sub_S = mk_in [inter_S] [set_F] typ_aF; val lhs = Inf_const (HOLogic.mk_setT typ_aF) $ (mk_image (absfree (dest_Free var_A) (Image_const typ_aF $ rel_pairs $ setF_sub_A)) $ var_S); val rhs = Image_const typ_aF $ rel_pairs $ setF_sub_S; val concl = HOLogic.mk_Trueprop (mk_leq lhs rhs); in Logic.all var_S (Logic.list_implies (prems, concl)) end; val rel_Inters = map2 rel_Inter_from_set_F (var_As ~~ var_Ss) sets_F; (* map_F_Just = map_F Just *) val map_F_Just = let val option_tys = map mk_MaybeT alphas; val somes = map Just_const alphas; in list_comb (subst_atomic_types (betas ~~ option_tys) map_F, somes) end; fun mk_set_F'_tm typ_a set_F = let val typ_aset = HOLogic.mk_setT typ_a; (* set_F' x = (\y\{y. REL (map_F Just x) y}. UNION (set_F y) set_Maybe) *) val sbind = mk_UNION (subst_Maybe alphas set_F $ var_mx) (set_Maybe_const typ_a); val collection = HOLogic.Collect_const typ_MaybeF $ absfree (dest_Free var_mx) (subst_Maybe alphas equiv_rel_a $ (map_F_Just $ var_x) $ var_mx); val set_F'_tm = lambda var_x (Inf_const typ_aset $ (mk_image (absfree (dest_Free var_mx) sbind) $ collection)); in set_F'_tm end val sets = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val sets' = map2 mk_set_F'_tm alphas sets; val (Iwits, wit_goals) = prepare_wits true repT wits opts alphas wits_F var_as var_as' sets' lthy; val goals = rel_pos_distr :: rel_Inters @ (case wits of NONE => [] | _ => wit_goals); val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; fun after_qed thmss lthy = (case thmss of [rel_pos_distr_thm0] :: thmss => let val equiv_thm' = REL_rewr_all lthy equiv_thm; val rel_pos_distr_thm = @{thm equivp_add_relconj} OF [equiv_thm', equiv_thm', rel_pos_distr_thm0]; val (rel_Inter_thms, wit_thmss) = apply2 (fn f => flat (f live thmss)) (take, drop); (* construct map_G, sets_G, bd_G, rel_G and wits_G *) (* map_G f = abs_G o map_F f o rep_G *) val map_G = fold_rev lambda var_fs (HOLogic.mk_comp (HOLogic.mk_comp (subst_Ts abs_G betas, list_comb (map_F, var_fs)), rep_G)); val map_raw = fold_rev lambda var_fs (list_comb (map_F, var_fs)) |> subst_atomic_types (betas ~~ gammas); (* Define set_G and the three auxiliary definitions (set_F', F_in, F_in') *) fun mk_set_G var_A set_F lthy = let val typ_a = HOLogic.dest_setT (fastype_of var_A); val set_F'_tm = mk_set_F'_tm typ_a set_F val (set_F', lthy) = define lthy (suffix set_F "'") 1 set_F'_tm; (* set_G = set_F' o rep_G *) val set_G = HOLogic.mk_comp (#tm set_F', rep_G); (* F_in A = {x. set_F x \ A} *) val F_in_tm = lambda var_A (mk_in [var_A] [set_F] typ_aF); val (F_in, lthy) = define lthy (suffix set_F "_in") 1 F_in_tm; (* F_in' A = map_F Inr -` ({(x, y). REL x y} `` F_in (insert (Inl ()) (Inr ` A))) *) val F_in' = lambda var_A (mk_vimage map_F_Just (Image_const typ_MaybeF $ subst_Maybe alphas rel_pairs $ (subst_Maybe alphas (#tm F_in) $ mk_insert (mk_Nothing typ_a) (mk_image (Just_const typ_a) $ var_A)))); val (F_in', lthy) = define lthy (suffix set_F "_in'") 1 F_in'; in ((set_G, {set_F'=set_F', F_in=F_in, F_in'=F_in'}), lthy) end; val ((sets_G, set_F'_aux_defs), lthy) = @{fold_map 2} mk_set_G var_As sets_F lthy |>> split_list; (* bd_G = bd_F *) val bd_G = mk_bd_of_bnf deads alphas bnf_F; (* rel_F' A = BNF_Def.vimage2p (map_F Just) (map_F Just) ((\) OO rel_F (rel_Maybe A) OO (\)) *) val rel_Maybes = @{map 3} (fn v => fn aT => fn bT => rel_Maybe_const aT bT $ v); val rel_F'_tm = let val equiv_equiv_rel_option = subst_Ts equiv_rel_a' o map mk_MaybeT in mk_vimage2p map_F_Just (subst_atomic_types (alphas ~~ betas) map_F_Just) $ mk_OO [equiv_equiv_rel_option alphas, list_comb (rel_F_option, rel_Maybes var_Ps alphas betas)] (equiv_equiv_rel_option betas) end; val (rel_F', lthy) = define lthy (suffix rel_F_ab "'") (live+2) (fold_rev lambda var_Ps rel_F'_tm); (* rel_G A = vimage2p rep_G rep_G (rel_F' A) *) val rel_G = fold_rev lambda var_Ps (mk_vimage2p rep_G (subst_Ts rep_G betas) $ rel_F'_tm); val rel_raw = fold_rev lambda var_Ps rel_F'_tm |> subst_atomic_types (betas ~~ gammas); (* auxiliary lemmas *) val bd_card_order = bd_card_order_of_bnf bnf_F; val bd_cinfinite = bd_cinfinite_of_bnf bnf_F; val in_rel = in_rel_of_bnf bnf_F; val map_F_comp = map_comp_of_bnf bnf_F; val map_F_comp0 = map_comp0_of_bnf bnf_F; val map_F_cong = map_cong_of_bnf bnf_F; val map_F_id0 = map_id0_of_bnf bnf_F; val map_F_id = map_id_of_bnf bnf_F; val rel_conversep = rel_conversep_of_bnf bnf_F; val rel_flip = rel_flip_of_bnf bnf_F; val rel_eq_onp = rel_eq_onp_of_bnf bnf_F; val rel_Grp = rel_Grp_of_bnf bnf_F; val rel_OO = rel_OO_of_bnf bnf_F; val rel_map = rel_map_of_bnf bnf_F; val rel_refl_strong = rel_refl_strong_of_bnf bnf_F; val set_bd_thms = set_bd_of_bnf bnf_F; val set_map_thms = set_map_of_bnf bnf_F; (* map_F_respect: @{term "((=) ===> REL ===> REL) map_F map_F"} *) val map_F_respect = HOLogic.mk_Trueprop (fold_rev mk_rel_fun (map2 (fn xT => fn yT => HOLogic.eq_const (xT --> yT)) alphas betas @ [equiv_rel_a]) (equiv_rel_b) $ map_F $ map_F); fun map_F_respect_tac ctxt = HEADGOAL (EVERY' [REPEAT_DETERM_N (live + 1) o rtac ctxt @{thm rel_funI}, hyp_subst_tac ctxt, rtac ctxt (BNF_FP_Util.split_conj_prems live rel_pos_distr_thm0 OF replicate live @{thm Grp_conversep_nonempty} RS rev_mp), rtac ctxt impI, dtac ctxt @{thm predicate2D}, etac ctxt @{thm relcomppI2[rotated]}, rtac ctxt (rel_Grp RS @{thm predicate2_eqD} RS iffD2), rtac ctxt @{thm GrpI[OF refl]}, REPEAT_DETERM o resolve_tac ctxt [CollectI, conjI, subset_UNIV], rtac ctxt (rel_flip RS iffD2), rtac ctxt (rel_Grp RS @{thm predicate2_eqD} RS iffD2), rtac ctxt @{thm GrpI[OF refl]}, REPEAT_DETERM o resolve_tac ctxt [CollectI, conjI, subset_UNIV], SELECT_GOAL (unfold_thms_tac ctxt (rel_eq_onp :: @{thms Grp_conversep_eq_onp})), etac ctxt @{thm predicate2D[OF rel_conj_eq_onp, rotated]}, rtac ctxt equiv_thm']); val map_F_respect_thm = prove lthy [] map_F_respect map_F_respect_tac; val rel_funD = mk_rel_funDN (live+1); val map_F_rsp = (rel_funD map_F_respect_thm) OF (replicate live refl); fun map_F_rsp_of tms ctxt = (infer_instantiate' ctxt (NONE :: NONE :: map (SOME o Thm.cterm_of ctxt) tms) map_F_rsp) val qthms = let fun equivp_THEN thm = REL_rewr_all lthy equiv_thm RS thm; fun Quotient3_THEN thm = REL_rewr_all lthy quot_thm RS thm; in {abs_rep = Quotient3_THEN @{thm Quotient3_abs_rep}, rel_abs = Quotient3_THEN @{thm Quotient3_rel_abs}, rep_abs = Quotient3_THEN @{thm Quotient3_rep_abs}, rep_reflp = Quotient3_THEN @{thm Quotient3_rep_reflp}, rep_abs_rsp = Quotient3_THEN @{thm rep_abs_rsp}, reflp = equivp_THEN @{thm equivp_reflp}, symp = equivp_THEN @{thm equivp_symp}, transp = equivp_THEN @{thm equivp_transp}, REL = REL_def} end; (* lemma REL_OO_REL_left: REL OO REL OO R = REL OO R *) val REL_OO_REL_left_thm = let val tm = mk_Trueprop_eq (mk_OO [equiv_rel_a, equiv_rel_a] var_R, mk_relcompp equiv_rel_a var_R) fun tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, rtac ctxt ext, rtac ctxt iffI, TWICE (etac ctxt @{thm relcomppE}), rtac ctxt @{thm relcomppI}, etac ctxt (#transp qthms), TWICE (assume_tac ctxt), etac ctxt @{thm relcomppE}, etac ctxt @{thm relcomppI}, rtac ctxt @{thm relcomppI}, rtac ctxt (#reflp qthms), assume_tac ctxt]); in prove lthy [var_R] tm tac end; (* Generate theorems related to the setters *) val map_F_fs = list_comb (map_F, var_fs); (* aset aset asetset bset typ_b typ_b *) fun mk_set_F'_thmss (((((var_A, var_A'), var_S), var_B), var_b), var_b') set_F {set_F', F_in, F_in'} rel_Inter_thm set_map_F_thm (idx, vf) = let val (var_f, var_fs') = case vf of (f :: fs) => (f, fs) | _ => error "won't happen"; val typ_a = fastype_of var_f |> dest_funT |> fst; val typ_b = fastype_of var_b; val (typ_asetset, typ_aset) = `HOLogic.mk_setT (fastype_of var_A); val map_F_fs_x = map_F_fs $ var_x; (* F_in'_mono: A \ B \ F_in' A \ F_in' B *) val F_in'_mono_tm = mk_Trueprop_implies ([mk_leq var_A var_A'], mk_leq (#tm F_in' $ var_A) (#tm F_in' $ var_A')); fun F_in'_mono_tac ctxt = unfold_thms_tac ctxt [#def F_in', #def F_in] THEN HEADGOAL (EVERY' [rtac ctxt subsetI, etac ctxt vimageE, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, hyp_subst_tac ctxt, rtac ctxt (vimageI OF [refl]), rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, assume_tac ctxt ORELSE' rtac ctxt refl, rtac ctxt CollectI, etac ctxt subset_trans, etac ctxt (@{thm insert_mono} OF [@{thm image_mono}])]); val F_in'_mono_thm = prove lthy [var_A, var_A'] F_in'_mono_tm F_in'_mono_tac; (* F_in'_Inter: F_in' (\S) = (\A\S. F_in' A) *) val F_in'_Inter_tm = mk_Trueprop_eq ((#tm F_in' $ (Inf_const typ_aset $ var_S)), (Inf_const (HOLogic.mk_setT typ_aF) $ (mk_image (#tm F_in') $ var_S))); fun F_in'_Inter_tac ctxt = Local_Defs.unfold_tac ctxt [#def F_in', #def F_in] THEN HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (HOLogic.mk_eq (var_S, bot_const typ_asetset)))] @{thm case_split}) THEN' EVERY' [ hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms Inter_empty INT_empty UNIV_sum_unit_conv}), rtac ctxt @{thm set_eqI}, rtac ctxt iffI, rtac ctxt UNIV_I, rtac ctxt (vimageI OF [refl]), rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, rtac ctxt (#reflp qthms), rtac ctxt CollectI, rtac ctxt subset_UNIV, etac ctxt @{thm exE[OF ex_in_conv[THEN iffD2]]}, EqSubst.eqsubst_tac ctxt [0] @{thms image_INT[of _ UNIV _ id, simplified]}, rtac ctxt @{thm inj_Inr}, assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms INT_extend_simps vimage_INT[symmetric]}), rtac ctxt @{thm arg_cong2[where f=vimage, OF refl]}, rtac ctxt equalityI, rtac ctxt subsetI, rtac ctxt @{thm InterI}, etac ctxt imageE, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, hyp_subst_tac ctxt, rtac ctxt @{thm ImageI[OF CollectI]}, etac ctxt @{thm case_prodI} ORELSE' (SELECT_GOAL (unfold_thms_tac ctxt @{thms prod.case}) THEN' rtac ctxt @{thm refl}), rtac ctxt CollectI, etac ctxt subset_trans, etac ctxt @{thm INT_lower[OF imageI]}, rtac ctxt (@{thm subset_trans} OF [asm_rl, rel_Inter_thm]), K (unfold_thms_tac ctxt @{thms image_image}), rtac ctxt subset_refl, K (unfold_thms_tac ctxt @{thms INT_extend_simps ex_in_conv[symmetric]}), rtac ctxt exI, rtac ctxt imageI, assume_tac ctxt, rtac ctxt exI, rtac ctxt @{thm InterI}, etac ctxt imageE, hyp_subst_tac ctxt, rtac ctxt @{thm insertI1}]); val F_in'_Inter_thm = prove lthy [var_S] F_in'_Inter_tm F_in'_Inter_tac; (* set_F'_respect: (REL ===> (=)) set_F' set_F' *) val set_F'_respect_tm = HOLogic.mk_Trueprop (mk_rel_fun equiv_rel_a (HOLogic.eq_const typ_aset) $ #tm set_F' $ #tm set_F'); fun set_F'_respect_tac ctxt = unfold_thms_tac ctxt (#def set_F' :: @{thms rel_fun_def}) THEN HEADGOAL (EVERY' [TWICE (rtac ctxt allI), rtac ctxt impI, dtac ctxt (map_F_rsp_of (map Just_const alphas) ctxt), rtac ctxt @{thm INF_cong}, rtac ctxt @{thm Collect_eqI}, rtac ctxt iffI, etac ctxt (#transp qthms OF [#symp qthms]), assume_tac ctxt, etac ctxt (#transp qthms), assume_tac ctxt, rtac ctxt refl]); (* F_in'_alt2: F_in' A = {x. set_F' x \ A} *) val F_in'_alt2_tm = mk_Trueprop_eq (#tm F_in' $ var_A, mk_in [var_A] [#tm set_F'] typ_aF); fun F_in'_alt2_tac ctxt = HEADGOAL (rtac ctxt equalityI THEN' (Subgoal.FOCUS o K) (unfold_thms_tac ctxt (map #def [set_F', F_in', F_in])) ctxt THEN' EVERY' [rtac ctxt subsetI, rtac ctxt CollectI, rtac ctxt subsetI, dtac ctxt vimageD, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, dtac ctxt @{thm InterD}, rtac ctxt @{thm imageI[OF CollectI]}, etac ctxt (#symp qthms), etac ctxt @{thm UnionE}, etac ctxt imageE, hyp_subst_tac ctxt, etac ctxt @{thm subset_lift_sum_unitD}, etac ctxt @{thm setr.cases}, hyp_subst_tac ctxt, assume_tac ctxt]) THEN unfold_thms_tac ctxt [#def set_F'] THEN (HEADGOAL o EVERY') [rtac ctxt subsetI, etac ctxt CollectE, etac ctxt (subsetD OF [F_in'_mono_thm]), EqSubst.eqsubst_tac ctxt [0] [F_in'_Inter_thm], rtac ctxt @{thm InterI}] THEN REPEAT_DETERM (HEADGOAL (etac ctxt @{thm imageE} THEN' hyp_subst_tac ctxt)) THEN (HEADGOAL o EVERY') [etac ctxt CollectE, SELECT_GOAL (unfold_thms_tac ctxt (map #def [F_in', F_in])), rtac ctxt @{thm vimageI[OF refl]}, rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, etac ctxt (#symp qthms), rtac ctxt CollectI, rtac ctxt subsetI, rtac ctxt @{thm sum_insert_Inl_unit}, assume_tac ctxt, hyp_subst_tac ctxt, rtac ctxt imageI, rtac ctxt @{thm UnionI}, rtac ctxt imageI, assume_tac ctxt, rtac ctxt @{thm setr.intros[OF refl]}]; val F_in'_alt2_thm = prove lthy [var_A] F_in'_alt2_tm F_in'_alt2_tac; (* set_F'_alt: set_F' x = \{A. x \ F_in' A} *) val set_F'_alt_tm = mk_Trueprop_eq (#tm set_F' $ var_x, Inf_const typ_aset $ mk_Collect (var_A, HOLogic.mk_mem (var_x, #tm F_in' $ var_A))); fun set_F'_alt_tac ctxt = unfold_thms_tac ctxt [F_in'_alt2_thm] THEN HEADGOAL (EVERY' [rtac ctxt @{thm set_eqI}, rtac ctxt iffI, rtac ctxt @{thm InterI}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt subsetD, assume_tac ctxt, assume_tac ctxt, etac ctxt @{thm InterD}, rtac ctxt CollectI, rtac ctxt CollectI, rtac ctxt subset_refl]); val set_F'_alt_thm = prove lthy [var_x] set_F'_alt_tm set_F'_alt_tac; (* map_F_in_F_in'I: x \ F_in' B \ map_F f x \ F_in' (f ` B) *) val map_F_in_F_in'I_tm = mk_Trueprop_implies ([HOLogic.mk_mem (var_x, #tm F_in' $ var_A')], HOLogic.mk_mem (map_F_fs_x, subst_b (#tm F_in') $ (mk_image var_f $ var_A'))); fun map_F_in_F_in'I_tac ctxt = Local_Defs.unfold_tac ctxt ([#def F_in', #def F_in, Bex_def] @ @{thms vimage_def Image_iff}) THEN HEADGOAL (EVERY' [etac ctxt @{thm CollectE}, etac ctxt exE, etac ctxt conjE, etac ctxt @{thm CollectE}, etac ctxt @{thm CollectE}, dtac ctxt @{thm case_prodD}, rtac ctxt @{thm CollectI}, rtac ctxt exI, rtac ctxt @{thm conjI[rotated]}, rtac ctxt @{thm CollectI}, rtac ctxt @{thm case_prodI}, dtac ctxt (map_F_rsp_of (map mk_Maybe_map var_fs) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms o_def map_sum.simps})), assume_tac ctxt, rtac ctxt CollectI, SELECT_GOAL (unfold_thms_tac ctxt set_map_thms), etac ctxt @{thm image_map_sum_unit_subset}]); val map_F_in_F_in'I_thm = prove lthy (var_A' :: var_x :: var_fs) map_F_in_F_in'I_tm map_F_in_F_in'I_tac; (* REL_preimage_eq: C \ range f \ {} \ {(a, b). REL a b} `` {x. set_F x \ f -` C} = map_F f -` {(a, b). REL a b} `` {x. set_F x \ C} *) val REL_preimage_eq_tm = mk_Trueprop_implies ([HOLogic.mk_not (HOLogic.mk_eq (HOLogic.mk_binop @{const_name inf} (var_B, mk_image var_f $ HOLogic.mk_UNIV typ_a), bot_const (HOLogic.mk_setT typ_b)))], HOLogic.mk_eq (Image_const typ_aF $ rel_pairs $ mk_in [mk_vimage var_f var_B] [set_F] typ_aF, mk_vimage map_F_fs (Image_const typ_bF $ subst_b rel_pairs $ mk_in [var_B] [subst_b set_F] typ_bF))); (* Bs \ range fs \ {} \ set_F xb \ Bs \ REL xb (map_F fs x) \ x \ {(x, x'). REL x x'} `` {x. set_F x \ fs -` Bs} *) fun subgoal_tac {context = ctxt, params, ...} = let val (x, y) = case params of [(_, x), _, (_, y)] => (x, y) | _ => error "won't happen"; val cond = HOLogic.mk_conj (apply2 HOLogic.mk_mem ((var_b, var_B), (var_b', var_B))); (* ["\x y. x \ B \ y \ B", "(Grp UNIV f_1)\\"] *) val cvars = var_fs |> maps (fn f => let val fT = fastype_of f in map (SOME o Thm.cterm_of ctxt) [if f = var_f then fold_rev lambda [var_b, var_b'] cond else HOLogic.eq_const (range_type fT), mk_conversep (mk_Grp (HOLogic.mk_UNIV (domain_type fT)) f)] end); val rel_pos_distr_thm_inst = infer_instantiate' ctxt (cvars @ [SOME y,SOME x]) (@{thm predicate2D} OF [rel_pos_distr_thm]); (* GrpI[of "map_F f1 .. fN" x, OF refl CollectI, OF "B1 \ UNIV \ ... \ Bn \ UNIV"] *) fun subset_UNIVs n = fold (fn a => fn b => conjI OF [a, b]) (replicate (n-1) @{thm subset_UNIV}) @{thm subset_UNIV}; val GrpI_inst = infer_instantiate' ctxt (map SOME [Thm.cterm_of ctxt map_F_fs, x]) @{thm GrpI} OF [refl, CollectI] OF [subset_UNIVs live]; in EVERY [ HEADGOAL (Method.insert_tac ctxt [rel_pos_distr_thm_inst]), unfold_thms_tac ctxt [rel_conversep, rel_OO, rel_Grp], HEADGOAL (etac ctxt @{thm meta_impE}), REPEAT_DETERM_N (live-1) (HEADGOAL (rtac ctxt @{thm conjI[rotated]})), REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm relcompp_mem_Grp_neq_bot} ORELSE' rtac ctxt @{thm relcompp_eq_Grp_neq_bot})), HEADGOAL (EVERY' [etac ctxt @{thm meta_impE}, rtac ctxt @{thm relcomppI}, rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI}, rtac ctxt rel_refl_strong]), REPEAT_DETERM_N idx (HEADGOAL (rtac ctxt refl)), HEADGOAL (rtac ctxt conjI THEN' TWICE (etac ctxt subsetD THEN' assume_tac ctxt)), REPEAT_DETERM_N (live-idx-1) (HEADGOAL (rtac ctxt refl)), HEADGOAL (EVERY' [rtac ctxt @{thm relcomppI}, assume_tac ctxt, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, rtac ctxt GrpI_inst, rtac ctxt (#reflp qthms), etac ctxt @{thm relcomppE}, etac ctxt @{thm relcomppE}, etac ctxt @{thm relcomppE}, etac ctxt @{thm conversepE}, etac ctxt @{thm GrpE}, hyp_subst_tac ctxt, rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, assume_tac ctxt, EqSubst.eqsubst_asm_tac ctxt [1] rel_map, EqSubst.eqsubst_asm_tac ctxt [1] [in_rel_of_bnf bnf_F], etac ctxt exE, etac ctxt CollectE, etac ctxt conjE, etac ctxt conjE, etac ctxt CollectE, hyp_subst_tac ctxt, rtac ctxt CollectI]), unfold_thms_tac ctxt set_map_thms, HEADGOAL (rtac ctxt (subsetI OF [vimageI] OF [refl]) THEN' etac ctxt @{thm imageE} THEN' hyp_subst_tac ctxt), REPEAT_DETERM_N 6 (HEADGOAL (etac ctxt Drule.thin_rl)), REPEAT_DETERM_N (live-1) (HEADGOAL (etac ctxt conjE)), HEADGOAL (EVERY' [dtac ctxt subsetD, assume_tac ctxt, etac ctxt CollectE]), unfold_thms_tac ctxt @{thms split_beta}, HEADGOAL (etac ctxt conjunct2)] end; fun REL_preimage_eq_tac ctxt = HEADGOAL (EVERY' [rtac ctxt @{thm set_eqI}, rtac ctxt iffI, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, rtac ctxt (vimageI OF [refl]), rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, etac ctxt map_F_rsp, rtac ctxt CollectI, EqSubst.eqsubst_tac ctxt [0] [set_map_F_thm], etac ctxt @{thm subset_vimage_image_subset}, etac ctxt vimageE, etac ctxt @{thm ImageE}, TWICE (etac ctxt CollectE), dtac ctxt @{thm case_prodD}, hyp_subst_tac ctxt, Subgoal.FOCUS_PARAMS subgoal_tac ctxt]); val REL_preimage_eq_thm = prove lthy (var_B :: var_fs) REL_preimage_eq_tm REL_preimage_eq_tac; (* set_map_F': set_F' (map_F f x) = f ` set_F' x *) val set_map_F'_tm = mk_Trueprop_eq (subst_b (#tm set_F') $ map_F_fs_x, mk_image var_f $ (#tm set_F' $ var_x)); fun set_map_F'_tac ctxt = HEADGOAL (EVERY' [rtac ctxt @{thm set_eqI}, rtac ctxt iffI, EqSubst.eqsubst_asm_tac ctxt [0] [set_F'_alt_thm], etac ctxt @{thm InterD}, rtac ctxt CollectI, rtac ctxt map_F_in_F_in'I_thm, SELECT_GOAL (unfold_thms_tac ctxt [F_in'_alt2_thm]), rtac ctxt CollectI, rtac ctxt subset_refl, SELECT_GOAL (unfold_thms_tac ctxt [set_F'_alt_thm]), rtac ctxt @{thm InterI}, etac ctxt imageE, etac ctxt CollectE, hyp_subst_tac ctxt, etac ctxt @{thm vimageD[OF InterD]}, rtac ctxt CollectI]) THEN (* map_F f x \ F_in' X \ x \ F_in' (f -` X) *) HEADGOAL (Subgoal.FOCUS_PARAMS (fn {context = ctxt, params, ...} => let val X = nth params 1 |> snd |> Thm.term_of; val Inr_img = mk_image (Just_const (HOLogic.dest_setT (fastype_of X))) $ X; fun cvars_of ctxt = map (SOME o Thm.cterm_of ctxt); val cut_thm = infer_instantiate' ctxt (cvars_of ctxt [Inr_img, var_f]) @{thm insert_Inl_int_map_sum_unit}; val preimage_thm = infer_instantiate' ctxt (cvars_of ctxt (filter (fn f => var_f <> f) var_fs |> map mk_Maybe_map)) (cut_thm RS REL_preimage_eq_thm); in EVERY [ unfold_thms_tac ctxt (map #def [F_in', F_in] @ preimage_thm :: map_F_comp :: @{thms lift_sum_unit_vimage_commute vimage_comp o_def map_sum.simps}), unfold_thms_tac ctxt [@{thm o_def[symmetric]}, map_F_comp0], Local_Defs.fold_tac ctxt @{thms vimage_comp}, HEADGOAL (etac ctxt (vimageI OF [refl]))] end) ctxt); (* set_F'_subset: set_F' x \ set_F x *) val set_F'_subset_tm = HOLogic.mk_Trueprop (mk_leq (#tm set_F' $ var_x) (set_F $ var_x)); fun set_F'_subset_tac ctxt = let val int_e_thm = infer_instantiate' ctxt (replicate 3 NONE @ [SOME (Thm.cterm_of ctxt (map_F_Just $ var_x))]) @{thm INT_E}; in HEADGOAL (EVERY' [SELECT_GOAL (Local_Defs.unfold_tac ctxt [#def set_F']), rtac ctxt subsetI, etac ctxt int_e_thm, SELECT_GOAL (unfold_thms_tac ctxt [set_map_F_thm]), etac ctxt @{thm UN_E}, etac ctxt imageE, hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms sum_set_simps singleton_iff}), hyp_subst_tac ctxt, assume_tac ctxt, etac ctxt notE, rtac ctxt CollectI, rtac ctxt (#reflp qthms)]) end; in ({F_in'_mono = F_in'_mono_thm, F_in'_Inter = F_in'_Inter_thm, set_F'_respect = prove lthy [] set_F'_respect_tm set_F'_respect_tac, F_in'_alt2 = F_in'_alt2_thm, set_F'_alt = set_F'_alt_thm, map_F_in_F_in'I = map_F_in_F_in'I_thm, set_map_F' = prove lthy (var_x :: var_fs) set_map_F'_tm set_map_F'_tac, set_F'_subset = prove lthy [var_x] set_F'_subset_tm set_F'_subset_tac, set_F'_def = #def set_F', F_in_def = #def F_in, F_in'_def = #def F_in'}, (idx + 1, var_fs')) end; val set_F'_thmss = @{fold_map 5} mk_set_F'_thmss (var_As ~~ var_As' ~~ var_Ss ~~ var_Bs ~~ var_bs ~~ var_bs') sets_F set_F'_aux_defs rel_Inter_thms set_map_thms (0, var_fs) |> fst; (* F_in'D: x \ F_in' A \ \a\A. f a = g a \ REL (map_F f x) (map_F g x) *) fun rel_map_F_fs_map_F_gs subst fs gs = subst equiv_rel_b $ (list_comb (subst map_F, fs) $ var_x) $ (list_comb (subst map_F, gs) $ var_x); val F_in'D_thm = let fun mk_prem var_a var_Aset {F_in', ...} var_f var_g = [HOLogic.mk_mem (var_x, #tm F_in' $ var_Aset), mk_Ball var_Aset ((absfree (dest_Free var_a)) (HOLogic.mk_eq (var_f $ var_a, var_g $ var_a)))]; val prems = @{map 5} mk_prem var_as var_As set_F'_aux_defs var_fs' var_gs'; val F_in'D_tm = mk_Trueprop_implies (flat prems, rel_map_F_fs_map_F_gs (subst_Maybe betas) var_fs' var_gs'); fun map_F_rsp_of_case_sums fs ctxt = map_F_rsp_of (@{map 2} (fn f => fn T => BNF_FP_Util.mk_case_sum (Term.absdummy HOLogic.unitT (mk_Nothing T), f)) fs betas) ctxt; fun mk_var_fgs n = take n var_gs' @ drop n var_fs'; fun F_in'D_tac ctxt = EVERY (unfold_thms_tac ctxt (maps (fn {F_in'_def, F_in_def, ...} => [F_in'_def, F_in_def]) set_F'_thmss) :: map (REPEAT_DETERM_N live o HEADGOAL) [etac ctxt vimageE, etac ctxt @{thm ImageE}, etac ctxt CollectE THEN' etac ctxt CollectE, dtac ctxt @{thm case_prodD}] @ HEADGOAL (hyp_subst_tac ctxt THEN' rotate_tac (~live)) :: map (fn i => (HEADGOAL o EVERY') [if i < live then rtac ctxt (#transp qthms) else K all_tac, Ctr_Sugar_Tactics.select_prem_tac ctxt live (dresolve_tac ctxt [asm_rl]) i, rtac ctxt (#transp qthms), dtac ctxt (map_F_rsp_of_case_sums (mk_var_fgs (i-1)) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms case_sum_o_inj(2)})), etac ctxt (#symp qthms), dtac ctxt (map_F_rsp_of_case_sums (mk_var_fgs i) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms case_sum_o_inj(2)})), EqSubst.eqsubst_tac ctxt [0] [map_F_cong OF replicate i refl @ asm_rl :: replicate (live-i) refl], rtac ctxt @{thm sum.case_cong[OF refl refl]}, etac ctxt bspec, hyp_subst_tac ctxt, etac ctxt @{thm subset_lift_sum_unitD}, assume_tac ctxt, assume_tac ctxt]) (1 upto live)); in prove lthy (var_x :: var_As @ var_fs' @ var_gs') F_in'D_tm F_in'D_tac end; (* map_F_cong': (\a. a \ set_F' x \ f a = g a) \ REL (map_F f x) (map_F g x) *) val map_F_cong'_thm = let fun mk_prem {set_F', ...} var_a var_f var_g = Logic.all var_a (mk_Trueprop_implies ([HOLogic.mk_mem (var_a, #tm set_F' $ var_x)], HOLogic.mk_eq (var_f $ var_a, var_g $ var_a))); val map_F_cong'_tm = Logic.list_implies (@{map 4} mk_prem set_F'_aux_defs var_as var_fs var_gs, HOLogic.mk_Trueprop (rel_map_F_fs_map_F_gs I var_fs var_gs)); fun Just_o_fun bT f = HOLogic.mk_comp (Just_const bT, f); fun map_F_Just_o_funs fs = list_comb (subst_Maybe betas map_F, map2 Just_o_fun betas fs) $ var_x; fun map_F_cong'_tac ctxt = let val map_F_respect_inst = map_F_rsp |> infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) (map map_F_Just_o_funs [var_fs, var_gs] @ map fromJust_const betas)) |> Local_Defs.unfold ctxt (map_F_comp :: @{thms o_assoc Fun.o_apply[where f=projr and g=Inr, unfolded sum.sel] id_def[symmetric]}) |> Local_Defs.unfold ctxt @{thms id_comp}; in HEADGOAL (rtac ctxt map_F_respect_inst THEN' rtac ctxt F_in'D_thm) THEN EVERY (map (fn {F_in'_alt2, ...} => unfold_thms_tac ctxt [F_in'_alt2] THEN HEADGOAL (EVERY' [rtac ctxt CollectI, rtac ctxt subset_refl, rtac ctxt ballI, SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt @{thm arg_cong[where f=Inr]}, asm_full_simp_tac ctxt])) set_F'_thmss) end; in prove lthy (var_x :: var_fs @ var_gs) map_F_cong'_tm map_F_cong'_tac end; (* rel_F'_set: "rel_F' P x y \ (\z. set_F' z \ {(x, y). P x y} \ REL (map_F fst z) x \ REL (map_F snd z) y)" *) val rel_F'_set_thm = let val lhs = list_comb (#tm rel_F', var_Ps) $ var_x $ var_y; fun mk_subset_A var_a var_b var_P {set_F', ...} = let val collect_A = mk_case_prod (var_a, var_b) (var_P $ var_a $ var_b); in mk_leq (subst_atomic_types (alphas ~~ typ_pairs) (#tm set_F') $ var_z) collect_A end; val subset_As = @{map 4} mk_subset_A var_as var_bs var_Ps set_F'_aux_defs; fun mk_map mfs f z = Term.list_comb (mfs, map (fst o Term.strip_comb o f) var_ts) $ z; val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf_F; val map_F_snd = mk_map_of_bnf deads typ_pairs betas bnf_F; val map_fst = equiv_rel_a $ (mk_map map_F_fst HOLogic.mk_fst var_z) $ var_x; val map_snd = equiv_rel_b $ (mk_map map_F_snd HOLogic.mk_snd var_z) $ var_y; val rhs = let val (z, T) = dest_Free var_z in HOLogic.mk_exists (z, T, fold_rev (fn a => fn b => HOLogic.mk_conj (a, b)) (subset_As @ [map_fst]) map_snd) end; val rel_F'_set_tm = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); val maybePairsTs = map HOLogic.mk_prodT (map mk_MaybeT alphas ~~ map mk_MaybeT betas) fun mk_map_prod_projr aT bT = let val (mabT, (maT, mbT)) = `HOLogic.mk_prodT (apply2 mk_MaybeT (aT, bT)); val map_prod_const = Const (@{const_name map_prod}, (maT --> aT) --> (mbT --> bT) --> mabT --> HOLogic.mk_prodT (aT, bT)); in map_prod_const $ fromJust_const aT $ fromJust_const bT end; fun exI_OF_tac ctxt tm = rtac ctxt (infer_instantiate' ctxt (NONE :: [SOME (Thm.cterm_of ctxt tm)]) exI); (* REL (map_F Inr x) (map_F fst z) \ REL (map_F snd z) (map_F Inr y) \ set_F z \ {(x, y). rel_sum (=) P x y} \ \z. set_F' z \ {(x, y). P x y} \ REL (map_F fst z) x \ REL (map_F snd z) y *) fun subgoal1_tac {context = ctxt, params, ...} = let val z = (case params of (_ :: _ :: (_, ct) :: _) => Thm.term_of ct | _ => error "won't happen"); val map_F_projr_z = list_comb (mk_map_of_bnf deads maybePairsTs typ_pairs bnf_F, map2 mk_map_prod_projr alphas betas) $ z; in HEADGOAL (exI_OF_tac ctxt map_F_projr_z) THEN HEADGOAL (EVERY' (maps (fn {set_F'_subset, set_F'_respect, set_map_F', ...} => [rtac ctxt conjI, dtac ctxt (set_F'_subset RS @{thm order_trans}), TWICE (dtac ctxt (set_F'_respect RS @{thm rel_funD})), SELECT_GOAL (unfold_thms_tac ctxt [set_map_F']), etac ctxt @{thm in_rel_sum_in_image_projr}, TWICE (assume_tac ctxt)]) set_F'_thmss)) THEN HEADGOAL (EVERY' (map (fn Ts => FIRST' [dtac ctxt (map_F_rsp_of (map fromJust_const Ts) ctxt), etac ctxt sym , assume_tac ctxt]) [alphas, betas])) THEN unfold_thms_tac ctxt (map_F_comp :: @{thms fst_comp_map_prod snd_comp_map_prod comp_projr_Inr} @ [map_F_id]) THEN HEADGOAL (rtac ctxt conjI) THEN HEADGOAL (etac ctxt (#symp qthms) THEN' assume_tac ctxt ORELSE' (EVERY' (maps (fn Ts => [dtac ctxt (map_F_rsp_of (map fromJust_const Ts) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms fst_comp_map_prod snd_comp_map_prod comp_projr_Inr} @ [map_F_id])), assume_tac ctxt]) [alphas, betas]))) end; (* set_F' z \ {(x, y). P x y} \ REL (map_F fst z) x \ REL (map_F snd z) y \ \b. REL (map_F Inr x) b \ (\ba. rel_F (rel_sum (=) P) b ba \ REL ba (map_F Inr y)) *) fun subgoal2_tac {context = ctxt, params, ...} = let val z = (case params of ((_, ct) :: _) => Thm.term_of ct | _ => error "won't happen"); fun exI_map_Ifs_tac mk_proj Ts = exI_OF_tac ctxt (list_comb (mk_map_of_bnf deads typ_pairs (map mk_MaybeT Ts) bnf_F, @{map 3} (fn var_t => fn {set_F', ...} => fn T => lambda var_t (BNF_FP_Util.mk_If (HOLogic.mk_mem (var_t, subst_Ts (#tm set_F') typ_pairs $ z)) (mk_Just (mk_proj var_t)) (mk_Nothing T))) var_ts set_F'_aux_defs Ts) $ z) fun mk_REL_trans_map_F n = (rotate_prems n (#transp qthms) OF [rel_funD map_F_respect_thm] OF (replicate live refl @ [#symp qthms])); in HEADGOAL (EVERY' [exI_map_Ifs_tac HOLogic.mk_fst alphas, rtac ctxt conjI, etac ctxt (mk_REL_trans_map_F 0)]) THEN unfold_thms_tac ctxt [map_F_comp, @{thm o_def}] THEN HEADGOAL (rtac ctxt map_F_cong'_thm) THEN REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm if_P[symmetric]})) THEN HEADGOAL (EVERY' [exI_map_Ifs_tac HOLogic.mk_snd betas, rtac ctxt conjI]) THEN unfold_thms_tac ctxt rel_map THEN HEADGOAL (rtac ctxt rel_refl_strong) THEN REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm subset_rel_sumI})) THEN HEADGOAL (etac ctxt (mk_REL_trans_map_F 1 OF [#symp qthms])) THEN unfold_thms_tac ctxt [map_F_comp, @{thm o_def}] THEN HEADGOAL (rtac ctxt map_F_cong'_thm) THEN REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm if_P})) end; fun rel_F'_set_tac ctxt = EVERY ([unfold_thms_tac ctxt (#def rel_F' :: #REL qthms :: @{thms vimage2p_def relcompp_apply}), HEADGOAL (rtac ctxt iffI), (HEADGOAL o TWICE) (etac ctxt exE THEN' etac ctxt conjE), HEADGOAL (EVERY' [dtac ctxt (in_rel RS iffD1), etac ctxt exE, TWICE (etac ctxt conjE), etac ctxt CollectE, hyp_subst_tac ctxt]), (REPEAT_DETERM_N (live-1) o HEADGOAL) (etac ctxt conjE), HEADGOAL (Subgoal.FOCUS_PARAMS subgoal1_tac ctxt THEN' etac ctxt exE), (REPEAT_DETERM_N (live+1) o HEADGOAL) (etac ctxt conjE), HEADGOAL (Subgoal.FOCUS_PARAMS subgoal2_tac ctxt)]); in prove lthy (var_x :: var_y :: var_Ps) rel_F'_set_tm rel_F'_set_tac end; (* tactics *) (* map_G_id0: abs_G \ map_F id \ rep_G = id *) fun map_G_id0_tac ctxt = HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [@{thm fun_eq_iff}, o_apply, map_F_id0, id_apply, #abs_rep qthms]), rtac ctxt allI, rtac ctxt refl]); (* map_G (g \ f) = map_G g \ map_G f *) fun map_G_comp0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, rtac ctxt sym, SELECT_GOAL (unfold_thms_tac ctxt [o_apply, map_F_comp0]), rtac ctxt (#rel_abs qthms), rtac ctxt map_F_rsp, rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms)]); (* map_G_cong: (\z. z \ set_G x \ f z = g z) \ map_G f x = map_G g x *) fun map_G_cong_tac ctxt = EVERY [Local_Defs.fold_tac ctxt (map #set_F'_def set_F'_thmss), unfold_thms_tac ctxt [o_apply], HEADGOAL (rtac ctxt (#rel_abs qthms) THEN' rtac ctxt map_F_cong'_thm), REPEAT_DETERM_N live (HEADGOAL (asm_full_simp_tac ctxt))]; (* set_G_map0_G: set_G \ map_G f = f ` set_G *) fun mk_set_G_map0_G_tac thms ctxt = HEADGOAL (rtac ctxt ext) THEN EVERY [unfold_thms_tac ctxt [o_apply], Local_Defs.fold_tac ctxt [#set_F'_def thms]] THEN HEADGOAL (EVERY' (map (rtac ctxt) [trans OF [#set_map_F' thms RS sym, sym] RS sym, @{thm rel_funD} OF [#set_F'_respect thms], #rep_abs qthms, map_F_rsp, #rep_reflp qthms])); (* bd_card_order: card_order bd_F *) fun bd_card_order_tac ctxt = HEADGOAL (rtac ctxt bd_card_order); (* bd_cinfinite: BNF_Cardinal_Arithmetic.cinfinite bd_F *) fun bd_cinfinite_tac ctxt = HEADGOAL (rtac ctxt bd_cinfinite); (*target: ordLeq3 (card_of (set_F' (rep_G x_))) bd_F*) fun mk_set_G_bd_tac thms set_bd_thm ctxt = EVERY [Local_Defs.fold_tac ctxt [#set_F'_def thms], unfold_thms_tac ctxt [o_apply], HEADGOAL (rtac ctxt (@{thm ordLeq_transitive} OF [@{thm card_of_mono1} OF [#set_F'_subset thms], set_bd_thm]))]; (* rel_compp: rel_G R OO rel_G S \ rel_G (R OO S) *) fun rel_compp_tac ctxt = EVERY [unfold_thms_tac ctxt [#REL qthms], HEADGOAL (TWICE (rtac ctxt @{thm vimage2p_relcompp_mono})), (unfold_thms_tac ctxt (REL_OO_REL_left_thm :: @{thms relcompp_assoc})), (unfold_thms_tac ctxt [Local_Defs.unfold ctxt @{thms eq_OO} (infer_instantiate' ctxt [HOLogic.eq_const HOLogic.unitT |> Thm.cterm_of ctxt |> SOME] @{thm sum.rel_compp})]), HEADGOAL (rtac ctxt rel_pos_distr_thm), unfold_thms_tac ctxt @{thms fun_eq_iff bot_apply bot_bool_def not_all eq_False not_not OO_def}, REPEAT_DETERM (HEADGOAL (resolve_tac ctxt [exI, conjI, @{thm rel_sum.intros(1)}, refl]))]; (* rel_G R_ = (\x y. \z. set_G z \ {(x, y). R x y} \ map_G fst z = x \ map_G snd z = y) *) fun rel_compp_Grp_tac ctxt = let val _ = () in EVERY [Local_Defs.fold_tac ctxt (@{thm Grp_def} :: map #set_F'_def set_F'_thmss), unfold_thms_tac ctxt [o_apply, @{thm mem_Collect_eq}, @{thm OO_Grp_alt}, @{thm vimage2p_def}], Local_Defs.fold_tac ctxt [Local_Defs.unfold ctxt @{thms vimage2p_def} (#def rel_F')], unfold_thms_tac ctxt [rel_F'_set_thm], HEADGOAL (TWICE (rtac ctxt ext)), HEADGOAL (rtac ctxt iffI), REPEAT_DETERM (ALLGOALS (eresolve_tac ctxt [exE, conjE])), HEADGOAL (rtac ctxt exI), REPEAT_FIRST (resolve_tac ctxt [conjI]), HEADGOAL (EVERY' (maps (fn {set_F'_respect, ...} => [etac ctxt @{thm subset_trans[rotated]}, rtac ctxt equalityD1, rtac ctxt (@{thm rel_funD} OF [set_F'_respect]), rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms)]) set_F'_thmss)), (HEADGOAL o TWICE o EVERY') [rtac ctxt (trans OF [asm_rl, #abs_rep qthms]), rtac ctxt (#rel_abs qthms), etac ctxt (rotate_prems 1 (#transp qthms)), rtac ctxt map_F_rsp, rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms) ], HEADGOAL (rtac ctxt exI THEN' rtac ctxt conjI), (REPEAT_DETERM_N live o HEADGOAL o EVERY') [assume_tac ctxt, rtac ctxt conjI], (HEADGOAL o TWICE o EVERY') [ hyp_subst_tac ctxt, rtac ctxt (#rep_abs_rsp qthms), rtac ctxt map_F_rsp, rtac ctxt (#rep_reflp qthms)]] end; fun pred_G_set_G_tac ctxt = HEADGOAL (rtac ctxt refl); val tactics = map_G_id0_tac :: map_G_comp0_tac :: map_G_cong_tac :: map mk_set_G_map0_G_tac set_F'_thmss @ bd_card_order_tac :: bd_cinfinite_tac :: map2 mk_set_G_bd_tac set_F'_thmss set_bd_thms @ rel_compp_tac :: rel_compp_Grp_tac :: [pred_G_set_G_tac]; (* val wits_G = [abs_G o wit_F] *) val (wits_G, wit_thms) = let val wit_F_thmss = map (map2 (fn set_F' => fn wit => (#set_F'_subset set_F' RS set_mp RS wit) |> unfold_thms lthy [#set_F'_def set_F']) set_F'_thmss) (wit_thmss_of_bnf bnf_F); val (wits_F, wit_F_thmss) = split_list (filter_out (fn ((J, _), _) => exists (fn (I, _) => subset (op =) (I, J)) Iwits) (wits_F ~~ wit_F_thmss)); fun mk_wit (I, wit) = let val vars = (map (fn n => nth var_as n) I) in fold_rev lambda vars (abs_G $ list_comb (wit, vars)) end; in (map mk_wit (Iwits @ wits_F), wit_thmss @ flat wit_F_thmss) end; fun mk_wit_tacs ({set_F'_def, set_F'_respect, ...} :: set_F'_thmss) (w :: ws) ctxt = EVERY [unfold_thms_tac ctxt [@{thm o_def}, set_F'_respect RS @{thm rel_funD} OF [#rep_abs qthms OF [(#reflp qthms)]]], unfold_thms_tac ctxt [set_F'_def], HEADGOAL (etac ctxt w)] THEN mk_wit_tacs set_F'_thmss ws ctxt | mk_wit_tacs [] ws ctxt = mk_wit_tacs set_F'_thmss ws ctxt | mk_wit_tacs _ _ _ = all_tac; val (bnf_G, lthy) = bnf_def Dont_Inline (user_policy Note_Some) true I tactics (mk_wit_tacs [] wit_thms) NONE map_b rel_b pred_b set_bs (((((((Binding.empty, absT), map_G), sets_G), bd_G), wits_G), SOME rel_G), NONE) lthy; val old_defs = {sets = set_defs_of_bnf bnf_G, map = map_def_of_bnf bnf_G, rel = rel_def_of_bnf bnf_G}; val set_F'_defs = map (mk_abs_def o #set_F'_def) set_F'_thmss; val unfold_morphism = Morphism.thm_morphism "BNF" (unfold_thms lthy (defs @ #def REL :: set_F'_defs)); val (bnf_G, lthy) = morph_bnf_defs unfold_morphism bnf_G |> (fn bnf => note_bnf_defs bnf lthy); (* auxiliary lemmas transfer for transfer *) val rel_monoD_rotated = rotate_prems ~1 (rel_mono_of_bnf bnf_F RS @{thm predicate2D}); val REL_pos_distrI = let fun tac ctxt = EVERY [HEADGOAL (dtac ctxt (rotate_prems ~1 (rel_pos_distr_thm RS @{thm predicate2D}))), (REPEAT_DETERM o HEADGOAL) (rtac ctxt conjI ORELSE' assume_tac ctxt), (REPEAT_DETERM o HEADGOAL) (etac ctxt @{thm relcomppE}), HEADGOAL (dtac ctxt rel_monoD_rotated), (REPEAT_DETERM o HEADGOAL) (assume_tac ctxt ORELSE' rtac ctxt @{thm relcomppI})]; in prove lthy (var_x :: var_y' :: var_Ps @ var_Qs @ var_Rs) REL_pos_distrI_tm tac end; val rel_F_rel_F' = let val rel_F = mk_rel_of_bnf deads alphas betas bnf_F; val rel_F_rel_F'_tm = (rel_F, #tm rel_F') |> apply2 (fn R => HOLogic.mk_Trueprop (list_comb (R, var_Ps) $ var_x $ var_y)) |> Logic.mk_implies; fun rel_F_rel_F'_tac ctxt = EVERY [HEADGOAL (dtac ctxt (in_rel_of_bnf bnf_F RS iffD1)), unfold_thms_tac ctxt (rel_F'_set_thm :: @{thms mem_Collect_eq}), (REPEAT_DETERM o HEADGOAL) (eresolve_tac ctxt [exE, conjE]), HEADGOAL (rtac ctxt exI), HEADGOAL (EVERY' (maps (fn thms => [rtac ctxt conjI, rtac ctxt subsetI, dtac ctxt (set_mp OF [#set_F'_subset thms]), dtac ctxt subsetD, assume_tac ctxt, assume_tac ctxt]) set_F'_thmss)), (REPEAT_DETERM o HEADGOAL) (rtac ctxt conjI ORELSE' hyp_subst_tac ctxt THEN' rtac ctxt (#reflp qthms))] in prove lthy (var_x :: var_y :: var_Ps) rel_F_rel_F'_tm rel_F_rel_F'_tac end; fun inst_REL_pos_distrI n vs aTs bTs ctxt = infer_instantiate' ctxt (replicate n NONE @ (rel_Maybes vs aTs bTs |> map (SOME o Thm.cterm_of ctxt))) REL_pos_distrI; val Tss = {abs = typ_subst_atomic (alphas ~~ betas) absT, rep = repT, Ds0 = map TFree Ds0, deads = deads, alphas = alphas, betas = betas, gammas = gammas, deltas = deltas}; val thms = {map_F_rsp = map_F_rsp, rel_F'_def = #def rel_F', rel_F_rel_F' = rel_F_rel_F', rel_F'_set = rel_F'_set_thm, rel_monoD_rotated = rel_monoD_rotated} val transfer_consts = mk_quotient_transfer_tacs bnf_F Tss live qthms thms set_F'_thmss old_defs inst_REL_pos_distrI map_raw rel_raw (map (#tm o #set_F') set_F'_aux_defs); val quiet = exists (fn No_Warn_Transfer => true | _ => false) opts; in lthy |> BNF_Def.register_bnf plugins absT_name bnf_G |> mk_transfer_thms quiet bnf_F bnf_G absT_name transfer_consts (Quotient equiv_thm) Tss (defs @ #def REL :: set_F'_defs) end | _ => raise Match); in (goals, after_qed, #def REL :: defs, lthy) end; (** main commands **) local fun prepare_common prepare_name prepare_sort prepare_term prepare_thm - (((((plugins, raw_specs), raw_absT_name), raw_wits), xthm_opt), (map_b, rel_b, pred_b)) lthy = + (((((plugins, raw_specs), raw_absT_name), raw_wits), xthms_opt), (map_b, rel_b, pred_b)) lthy = let val absT_name = prepare_name lthy raw_absT_name; - val input_thm = - (case xthm_opt of - SOME xthm => prepare_thm lthy xthm - | NONE => Typedef.get_info lthy absT_name |> hd |> snd |> #type_definition); + fun bad_input input = + Pretty.chunks [Pretty.para ("Expected theorem(s) of either form:"), + Pretty.item [Pretty.enum "," "" "" [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T"}, + Syntax.pretty_term lthy @{term "reflp R"}]], + Pretty.item [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T"}], + Pretty.item [Syntax.pretty_term lthy @{term "type_definition Rep Abs A"}], + Pretty.para ("Got"), Pretty.enum "," "" "" (map (Thm.pretty_thm lthy) input)] + |> Pretty.string_of + |> error; + + fun no_refl qthm = + Pretty.chunks [Pretty.para ("Could not find a reflexivity rule for the Quotient theorem:"), + Pretty.item [Thm.pretty_thm lthy qthm], + Pretty.para ("Try supplying a reflexivity theorem manually or registering it in setup_lifting.")] + |> Pretty.string_of + |> error; + + fun find_equiv_thm_via_Quotient qthm = + let + val refl_thms = Lifting_Info.get_reflexivity_rules lthy + |> map (unfold_thms lthy @{thms reflp_eq[symmetric]}); + in + (case refl_thms RL [qthm RS @{thm Quotient_reflp_imp_equivp}] of + [] => no_refl qthm + | thm :: _ => thm) + end; + + val (lift_thm, equiv_thm) = + (case Option.map (prepare_thm lthy) xthms_opt of + SOME (thms as [qthm, _]) => + (case try (fn thms => @{thm Quotient_reflp_imp_equivp} OF thms) thms of + SOME equiv_thm => (qthm RS @{thm Quotient_Quotient3}, Quotient equiv_thm) + | NONE => bad_input thms) + | SOME [thm] => + (case try (fn thm => thm RS @{thm Quotient_Quotient3}) thm of + SOME qthm => (qthm, Quotient (find_equiv_thm_via_Quotient thm)) + | NONE => if can (fn thm => thm RS @{thm type_definition.Rep}) thm + then (thm, Typedef) + else bad_input [thm]) + | NONE => (case Lifting_Info.lookup_quotients lthy absT_name of + SOME {quot_thm = qthm, ...} => + let + val qthm = Thm.transfer' lthy qthm; + in + case [qthm] RL @{thms Quotient_eq_onp_typedef Quotient_eq_onp_type_copy} of + thm :: _ => (thm, Typedef) + | _ => (qthm RS @{thm Quotient_Quotient3}, + Quotient (find_equiv_thm_via_Quotient qthm)) + end + | NONE => (Typedef.get_info lthy absT_name |> hd |> snd |> #type_definition, Typedef)) + | SOME thms => bad_input thms); val wits = (Option.map o map) (prepare_term lthy) raw_wits; val specs = map (apsnd (apsnd (the_default @{sort type} o Option.map (prepare_sort lthy)))) raw_specs; - val which_bnf = (case (xthm_opt, Quotient_Info.lookup_quotients lthy absT_name) of - (NONE, SOME qs) => quotient_bnf (#equiv_thm qs, #quot_thm qs) - | (SOME _, _) => - if can (fn thm => thm RS @{thm bnf_lift_Quotient_equivp}) input_thm - then quotient_bnf (input_thm RS conjunct2, input_thm RS conjunct1 RS @{thm Quotient_Quotient3}) - else if can (fn thm => thm RS @{thm type_definition.Rep}) input_thm - then typedef_bnf - else Pretty.chunks [Pretty.para ("Expected theorem of either form:"), - Pretty.item [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T \ equivp R"}], - Pretty.item [Syntax.pretty_term lthy @{term "type_definition Rep Abs A"}], - Pretty.para ("Got"), Pretty.item [Thm.pretty_thm lthy input_thm]] - |> Pretty.string_of - |> error - | _ => typedef_bnf); - + val which_bnf = (case equiv_thm of + Quotient thm => quotient_bnf (thm, lift_thm) + | Typedef => typedef_bnf lift_thm); in - which_bnf input_thm wits specs map_b rel_b pred_b plugins lthy + which_bnf wits specs map_b rel_b pred_b plugins lthy end; fun prepare_lift_bnf prepare_name prepare_sort prepare_term prepare_thm = (fn (goals, after_qed, definitions, lthy) => lthy |> Proof.theorem NONE after_qed (map (single o rpair []) goals) |> Proof.refine_singleton (Method.Basic (fn ctxt => SIMPLE_METHOD (unfold_thms_tac ctxt definitions))) |> Proof.refine_singleton (Method.primitive_text (K I))) oo prepare_common prepare_name prepare_sort prepare_term prepare_thm o apfst (apfst (apsnd SOME)); fun prepare_solve prepare_name prepare_typ prepare_sort prepare_thm tacs = (fn (goals, after_qed, definitions, lthy) => lthy |> after_qed (map2 (fn goal => fn tac => [Goal.prove_sorry lthy [] [] goal (fn (ctxtprems as {context = ctxt, prems = _}) => unfold_thms_tac ctxt definitions THEN tac ctxtprems)]) goals (tacs (length goals)))) oo prepare_common prepare_name prepare_typ prepare_sort prepare_thm; in val lift_bnf_cmd = prepare_lift_bnf (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) - Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms); + Syntax.read_sort Syntax.read_term Attrib.eval_thms; fun lift_bnf args tacs = prepare_solve (K I) (K I) (K I) (K I) (K tacs) args; fun copy_bnf_tac {context = ctxt, prems = _} = REPEAT_DETERM (resolve_tac ctxt [bexI, conjI, UNIV_I, refl, subset_refl] 1); val copy_bnf = apfst (apfst (rpair NONE)) + #> apfst (apsnd (Option.map single)) #> prepare_solve (K I) (K I) (K I) (K I) (fn n => replicate n copy_bnf_tac); val copy_bnf_cmd = apfst (apfst (rpair NONE)) + #> apfst (apsnd (Option.map single)) #> prepare_solve (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) - Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms) + Syntax.read_sort Syntax.read_term Attrib.eval_thms (fn n => replicate n copy_bnf_tac); end; (** outer syntax **) local (* parsers *) val parse_wits = @{keyword "["} |-- (Parse.name --| @{keyword ":"} -- Scan.repeat Parse.term >> (fn ("wits", Ts) => Ts | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --| @{keyword "]"} || Scan.succeed []; fun parse_common_opts p = Scan.optional (@{keyword "("} |-- Parse.list1 (Parse.group (K "option") (Scan.first (p :: [Plugin_Name.parse_filter >> Plugins_Option, Parse.reserved "no_warn_transfer" >> K No_Warn_Transfer]))) --| @{keyword ")"}) []; val parse_lift_opts = Parse.reserved "no_warn_wits" >> K No_Warn_Wits |> parse_common_opts; val parse_copy_opts = parse_common_opts Scan.fail; val parse_xthm = Scan.option (Parse.reserved "via" |-- Parse.thm); +val parse_xthms = Scan.option (Parse.reserved "via" |-- Parse.thms1); in (* exposed commands *) val _ = Outer_Syntax.local_theory_to_proof @{command_keyword lift_bnf} "register a quotient type/subtype of a bounded natural functor (BNF) as a BNF" ((parse_lift_opts -- parse_type_args_named_constrained -- Parse.type_const -- parse_wits -- - parse_xthm -- parse_map_rel_pred_bindings) >> lift_bnf_cmd); + parse_xthms -- parse_map_rel_pred_bindings) >> lift_bnf_cmd); val _ = Outer_Syntax.local_theory @{command_keyword copy_bnf} "register a type copy of a bounded natural functor (BNF) as a BNF" ((parse_copy_opts -- parse_type_args_named_constrained -- Parse.type_const -- parse_xthm -- parse_map_rel_pred_bindings) >> copy_bnf_cmd); end; end;