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,3677 +1,3679 @@ (* 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.\ +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}? ; @{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 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/Cyclic_List.thy b/src/HOL/Datatype_Examples/Cyclic_List.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Datatype_Examples/Cyclic_List.thy @@ -0,0 +1,43 @@ +theory Cyclic_List imports + "HOL-Library.Confluent_Quotient" +begin + +inductive cyclic1 :: "'a list \ 'a list \ bool" for xs where + "cyclic1 xs (rotate1 xs)" + +abbreviation cyclic :: "'a list \ 'a list \ bool" where + "cyclic \ equivclp cyclic1" + +lemma cyclic_mapI: "cyclic xs ys \ cyclic (map f xs) (map f ys)" + by(induction rule: equivclp_induct) + (auto 4 4 elim!: cyclic1.cases simp add: rotate1_map[symmetric] intro: equivclp_into_equivclp cyclic1.intros) + +quotient_type 'a cyclic_list = "'a list" / cyclic by simp + +lemma map_respect_cyclic: includes lifting_syntax shows + "((=) ===> cyclic ===> cyclic) map map" + by(auto simp add: rel_fun_def cyclic_mapI) + +lemma confluentp_cyclic1: "confluentp cyclic1" + by(intro strong_confluentp_imp_confluentp strong_confluentpI)(auto simp add: cyclic1.simps) + +lemma cyclic_set_eq: "cyclic xs ys \ set xs = set ys" + by(induction rule: converse_equivclp_induct)(simp_all add: cyclic1.simps, safe, simp_all) + +lemma retract_cyclic1: + assumes "cyclic1 (map f xs) ys" + shows "\zs. cyclic xs zs \ ys = map f zs" + using assms by(auto simp add: cyclic1.simps rotate1_map intro: cyclic1.intros symclpI1) + +lemma confluent_quotient_cyclic1: + "confluent_quotient cyclic1 cyclic cyclic cyclic cyclic cyclic (map fst) (map snd) (map fst) (map snd) list_all2 list_all2 list_all2 set set" + by(unfold_locales) + (auto dest: retract_cyclic1 cyclic_set_eq simp add: list.in_rel list.rel_compp map_respect_cyclic[THEN rel_funD, OF refl] confluentp_cyclic1 intro: rtranclp_mono[THEN predicate2D, OF symclp_greater]) + +lift_bnf 'a cyclic_list [wits: "[]"] + subgoal by(rule confluent_quotient.subdistributivity[OF confluent_quotient_cyclic1]) + subgoal by(force dest: cyclic_set_eq) + subgoal by(auto elim: allE[where x="[]"]) + done + +end diff --git a/src/HOL/Datatype_Examples/Dlist.thy b/src/HOL/Datatype_Examples/Dlist.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Datatype_Examples/Dlist.thy @@ -0,0 +1,82 @@ +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 [wits: "[]"] + 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]) + subgoal by(auto simp add: dlist_eq_def vimage2p_def) + done + +end + +end diff --git a/src/HOL/Datatype_Examples/Free_Idempotent_Monoid.thy b/src/HOL/Datatype_Examples/Free_Idempotent_Monoid.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Datatype_Examples/Free_Idempotent_Monoid.thy @@ -0,0 +1,265 @@ +theory Free_Idempotent_Monoid imports + "HOL-Library.Confluent_Quotient" +begin + +inductive cancel1 :: "'a list \ 'a list \ bool" +where cancel1: "xs \ [] \ cancel1 (gs @ xs @ xs @ gs') (gs @ xs @ gs')" + +abbreviation cancel where "cancel \ cancel1\<^sup>*\<^sup>*" + +lemma cancel1_append_same1: + assumes "cancel1 xs ys" + shows "cancel1 (zs @ xs) (zs @ ys)" +using assms +proof cases + case (cancel1 ys gs gs') + from \ys \ []\ have "cancel1 ((zs @ gs) @ ys @ ys @ gs') ((zs @ gs) @ ys @ gs')" .. + with cancel1 show ?thesis by simp +qed + +lemma cancel_append_same1: "cancel (zs @ xs) (zs @ ys)" if "cancel xs ys" + using that by induction(blast intro: rtranclp.rtrancl_into_rtrancl cancel1_append_same1)+ + +lemma cancel1_append_same2: "cancel1 xs ys \ cancel1 (xs @ zs) (ys @ zs)" +by(cases rule: cancel1.cases)(auto intro: cancel1.intros) + +lemma cancel_append_same2: "cancel (xs @ zs) (ys @ zs)" if "cancel xs ys" + using that by induction(blast intro: rtranclp.rtrancl_into_rtrancl cancel1_append_same2)+ + +lemma cancel1_same: + assumes "xs \ []" + shows "cancel1 (xs @ xs) xs" +proof - + have "cancel1 ([] @ xs @ xs @ []) ([] @ xs @ [])" using assms .. + thus ?thesis by simp +qed + +lemma cancel_same: "cancel (xs @ xs) xs" + by(cases "xs = []")(auto intro: cancel1_same) + +abbreviation (input) eq :: "'a list \ 'a list \ bool" +where "eq \ equivclp cancel1" + +lemma eq_sym: "eq x y \ eq y x" + by(rule equivclp_sym) + +lemma equivp_eq: "equivp eq" by simp + +lemma eq_append_same1: "eq xs' ys' \ eq (xs @ xs') (xs @ ys')" + by(induction rule: equivclp_induct)(auto intro: cancel1_append_same1 equivclp_into_equivclp) + +lemma append_eq_cong: "\eq xs ys; eq xs' ys'\ \ eq (xs @ xs') (ys @ ys')" + by(induction rule: equivclp_induct)(auto intro: eq_append_same1 equivclp_into_equivclp cancel1_append_same2) + + +quotient_type 'a fim = "'a list" / eq by simp + +instantiation fim :: (type) monoid_add begin +lift_definition zero_fim :: "'a fim" is "[]" . +lift_definition plus_fim :: "'a fim \ 'a fim \ 'a fim" is "(@)" by(rule append_eq_cong) +instance by(intro_classes; transfer; simp) +end + +lemma plus_idem_fim [simp]: fixes x :: "'a fim" shows "x + x = x" +proof transfer + fix xs :: "'a list" + show "eq (xs @ xs) xs" + proof(cases "xs = []") + case False thus ?thesis using cancel1_same[of xs] by(auto) + qed simp +qed + + +inductive pcancel1 :: "'a list \ 'a list \ bool" where + pcancel1: "pcancel1 (concat xss) (concat yss)" if "list_all2 (\xs ys. xs = ys \ xs = ys @ ys) xss yss" + +lemma cancel1_into_pcancel1: "pcancel1 xs ys" if "cancel1 xs ys" + using that +proof cases + case (cancel1 xs gs gs') + let ?xss = "[gs, xs @ xs, gs']" and ?yss = "[gs, xs, gs']" + have "pcancel1 (concat ?xss) (concat ?yss)" by(rule pcancel1.intros) simp + then show ?thesis using cancel1 by simp +qed + +lemma pcancel_into_cancel: "cancel1\<^sup>*\<^sup>* xs ys" if "pcancel1 xs ys" + using that +proof cases + case (pcancel1 xss yss) + from pcancel1(3) show ?thesis unfolding pcancel1(1-2) + apply induction + apply simp + apply(auto intro: cancel_append_same1) + apply(rule rtranclp_trans) + apply(subst append_assoc[symmetric]) + apply(rule cancel_append_same2) + apply(rule cancel_same) + apply(auto intro: cancel_append_same1) + done +qed + +lemma pcancel1_cancel1_confluent: + assumes "pcancel1 xs ys" + and "cancel1 zs ys" + shows "\us. cancel us xs \ pcancel1 us zs" +proof - + let ?P = "\xs ys. xs = ys \ xs = ys @ ys" + consider ass as2 bs1 bss bs2 cs1 css ass' as2bs1 bss' bs2cs1 css' + where "ys = concat ass @ as2 @ bs1 @ concat bss @ bs2 @ cs1 @ concat css" + and "xs = concat ass' @ as2bs1 @ concat bss' @ bs2cs1 @ concat css'" + and "zs = concat ass @ as2 @ bs1 @ concat bss @ bs2 @ bs1 @ concat bss @ bs2 @ cs1 @ concat css" + and "list_all2 ?P ass' ass" + and "list_all2 ?P bss' bss" + and "list_all2 ?P css' css" + and "?P as2bs1 (as2 @ bs1)" + and "?P bs2cs1 (bs2 @ cs1)" + | ass as2 bs cs1 css ass' css' + where "ys = concat ass @ as2 @ bs @ cs1 @ concat css" + and "xs = concat ass' @ as2 @ bs @ cs1 @ as2 @ bs @ cs1 @ concat css'" + and "zs = concat ass @ as2 @ bs @ bs @ cs1 @ concat css" + and "list_all2 ?P ass' ass" + and "list_all2 ?P css' css" + proof - + from assms obtain xss bs as cs yss + where xs: "xs = concat xss" and zs: "zs = as @ bs @ bs @ cs" + and xss: "list_all2 (\xs ys. xs = ys \ xs = ys @ ys) xss yss" + and ys: "ys = as @ bs @ cs" + and yss: "concat yss = as @ bs @ cs" + and bs: "bs \ []" + by(clarsimp simp add: pcancel1.simps cancel1.simps) + from yss bs obtain ass as2 BS bcss + where yss: "yss = ass @ (as2 @ BS) # bcss" + and as: "as = concat ass @ as2" + and eq: "bs @ cs = BS @ concat bcss" + by(auto simp add: concat_eq_append_conv split: if_split_asm) + define bcss' where "bcss' = (if bcss = [] then [[]] else bcss)" + have bcss': "bcss' \ []" by(simp add: bcss'_def) + from eq consider us where "bs = BS @ us" "concat bcss = us @ cs" "bcss \ []" | + "BS = bs @ cs" "bcss = []" | + us where "BS = bs @ us" "cs = us @ concat bcss" + by(cases "bcss = []")(auto simp add: append_eq_append_conv2) + then show thesis + proof cases + case 1 + from 1(2,3) obtain bss bs2 cs1 css + where "bcss = bss @ (bs2 @ cs1) # css" + and us: "us = concat bss @ bs2" + and cs: "cs = cs1 @ concat css" by(auto simp add: concat_eq_append_conv) + with 1 xs xss ys yss zs as that(1)[of ass as2 BS bss bs2 cs1 css] show ?thesis + by(clarsimp simp add: list_all2_append2 list_all2_Cons2) + next + case 2 + with xs xss ys yss zs as show ?thesis + using that(1)[of ass as2 bs "[]" "[]" "[]" "[cs]" _ "as2 @ bs" "[]" "[]" "[cs]"] + using that(2)[of ass as2 bs cs "[]"] + by(auto simp add: list_all2_append2 list_all2_Cons2) + next + case 3 + with xs xss ys yss zs as show ?thesis + using that(1)[of ass as2 "[]" "[bs]" "[]" "us" "bcss" _ "as2" "[bs]"] that(2)[of ass as2 bs us bcss] + by(auto simp add: list_all2_append2 list_all2_Cons2) + qed + qed + then show ?thesis + proof cases + case 1 + let ?us = "concat ass' @ as2bs1 @ concat bss' @ bs2 @ bs1 @ concat bss' @ bs2cs1 @ concat css'" + have "?us = concat (ass' @ [as2bs1] @ bss' @ [bs2 @ bs1] @ bss' @ [bs2cs1] @ css')" by simp + also have "pcancel1 \ (concat (ass @ [as2 @ bs1] @ bss @ [bs2 @ bs1] @ bss @ [bs2 @ cs1] @ css))" + by(rule pcancel1.intros)(use 1 in \simp add: list_all2_appendI\) + also have "\ = zs" using 1 by simp + also have "cancel ?us xs" + proof - + define as2b where "as2b = (if as2bs1 = as2 @ bs1 then as2 else as2 @ bs1 @ as2)" + have as2bs1: "as2bs1 = as2b @ bs1" using 1 by(auto simp add: as2b_def) + define b2cs where "b2cs = (if bs2cs1 = bs2 @ cs1 then cs1 else cs1 @ bs2 @ cs1)" + have bs2cs1: "bs2cs1 = bs2 @ b2cs" using 1 by(auto simp add: b2cs_def) + have "?us = (concat ass' @ as2b) @ ((bs1 @ concat bss' @ bs2) @ (bs1 @ concat bss' @ bs2)) @ b2cs @ concat css'" + by(simp add: as2bs1 bs2cs1) + also have "cancel \ ((concat ass' @ as2b) @ (bs1 @ concat bss' @ bs2) @ b2cs @ concat css')" + by(rule cancel_append_same1 cancel_append_same2 cancel_same)+ + also have "\ = xs" using 1 bs2cs1 as2bs1 by simp + finally show ?thesis . + qed + ultimately show ?thesis by blast + next + case 2 + let ?us = "concat ass' @ as2 @ bs @ bs @ cs1 @ as2 @ bs @ bs @ cs1 @ concat css'" + have "?us = concat (ass' @ [as2 @ bs @ bs @ cs1 @ as2 @ bs @ bs @ cs1] @ css')" by simp + also have "pcancel1 \ (concat (ass @ [as2 @ bs @ bs @ cs1] @ css))" + by(intro pcancel1.intros list_all2_appendI)(simp_all add: 2) + also have "\ = zs" using 2 by simp + also have "cancel ?us xs" + proof - + have "?us = (concat ass' @ as2) @ (bs @ bs) @ cs1 @ as2 @ bs @ bs @ cs1 @ concat css'" by simp + also have "cancel \ ((concat ass' @ as2) @ bs @ cs1 @ as2 @ bs @ bs @ cs1 @ concat css')" + by(rule cancel_append_same1 cancel_append_same2 cancel_same)+ + also have "\ = (concat ass' @ as2 @ bs @ cs1 @ as2) @ (bs @ bs) @ cs1 @ concat css'" by simp + also have "cancel \ ((concat ass' @ as2 @ bs @ cs1 @ as2) @ bs @ cs1 @ concat css')" + by(rule cancel_append_same1 cancel_append_same2 cancel_same)+ + also have "\ = xs" using 2 by simp + finally show ?thesis . + qed + ultimately show ?thesis by blast + qed +qed + +lemma pcancel1_cancel_confluent: + assumes "pcancel1 xs ys" + and "cancel zs ys" + shows "\us. cancel us xs \ pcancel1 us zs" + using assms(2,1) + by(induction arbitrary: xs)(fastforce elim!: rtranclp_trans dest: pcancel1_cancel1_confluent)+ + +lemma cancel1_semiconfluent: + assumes "cancel1 xs ys" + and "cancel zs ys" + shows "\us. cancel us xs \ cancel us zs" + using pcancel1_cancel_confluent[OF cancel1_into_pcancel1, OF assms] + by(blast intro: pcancel_into_cancel) + +lemma semiconfluentp_cancel1: "semiconfluentp cancel1\\" + by(auto simp add: semiconfluentp_def rtranclp_conversep dest: cancel1_semiconfluent) + +lemma retract_cancel1_aux: + assumes "cancel1 ys (map f xs)" + shows "\zs. cancel1 zs xs \ ys = map f zs" + using assms + by cases(fastforce simp add: map_eq_append_conv append_eq_map_conv intro: cancel1.intros) + +lemma retract_cancel1: + assumes "cancel1 ys (map f xs)" + shows "\zs. eq xs zs \ ys = map f zs" + using retract_cancel1_aux[OF assms] by(blast intro: symclpI) + +lemma cancel1_set_eq: + assumes "cancel1 ys xs" + shows "set ys = set xs" + using assms by cases auto + +lemma eq_set_eq: "set xs = set ys" if "eq xs ys" + using that by(induction)(auto dest!: cancel1_set_eq elim!: symclpE) + +context includes lifting_syntax begin +lemma map_respect_cancel1: "((=) ===> cancel1 ===> eq) map map" + by(auto 4 4 simp add: rel_fun_def cancel1.simps intro: symclpI cancel1.intros) + +lemma map_respect_eq: "((=) ===> eq ===> eq) map map" + apply(intro rel_funI; hypsubst) + subgoal for _ f x y + by(induction rule: equivclp_induct) + (auto dest: map_respect_cancel1[THEN rel_funD, THEN rel_funD, OF refl] intro: eq_sym equivclp_trans) + done +end + +lemma confluent_quotient_fim: + "confluent_quotient cancel1\\ eq eq eq eq eq (map fst) (map snd) (map fst) (map snd) list_all2 list_all2 list_all2 set set" + by(unfold_locales)(auto dest: retract_cancel1 eq_set_eq simp add: list.in_rel list.rel_compp map_respect_eq[THEN rel_funD, OF refl] semiconfluentp_imp_confluentp semiconfluentp_cancel1)+ + +lift_bnf 'a fim [wits: "[]"] + subgoal for A B by(rule confluent_quotient.subdistributivity[OF confluent_quotient_fim]) + subgoal by(force dest: eq_set_eq) + subgoal by(auto elim: allE[where x="[]"]) + done + +end diff --git a/src/HOL/Datatype_Examples/LDL.thy b/src/HOL/Datatype_Examples/LDL.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Datatype_Examples/LDL.thy @@ -0,0 +1,339 @@ +theory LDL imports + "HOL-Library.Confluent_Quotient" +begin + +datatype 'a rexp = Atom 'a | Alt "'a rexp" "'a rexp" | Conc "'a rexp" "'a rexp" | Star "'a rexp" + +inductive ACI (infix "~" 64) where + a1: "Alt (Alt r s) t ~ Alt r (Alt s t)" +| a2: "Alt r (Alt s t) ~ Alt (Alt r s) t" +| c: "Alt r s ~ Alt s r" +| i: "r ~ Alt r r" +| R: "r ~ r" +| A: "r ~ r' \ s ~ s' \ Alt r s ~ Alt r' s'" +| C: "r ~ r' \ s ~ s' \ Conc r s ~ Conc r' s'" +| S: "r ~ r' \ Star r ~ Star r'" + +declare ACI.intros[intro] + +abbreviation ACIcl (infix "~~" 64) where "(~~) \ equivclp (~)" + +lemma eq_set_preserves_inter: + fixes eq set + assumes "\r s. eq r s \ set r = set s" and "Ss \ {}" + shows "(\As\Ss. {(x, x'). eq x x'} `` {x. set x \ As}) \ {(x, x'). eq x x'} `` {x. set x \ \ Ss}" + using assms by (auto simp: subset_eq) metis + +lemma ACI_map_respects: + fixes f :: "'a \ 'b" and r s :: "'a rexp" + assumes "r ~ s" + shows "map_rexp f r ~ map_rexp f s" + using assms by (induct r s rule: ACI.induct) auto + +lemma ACIcl_map_respects: + fixes f :: "'a \ 'b" and r s :: "'a rexp" + assumes "r ~~ s" + shows "map_rexp f r ~~ map_rexp f s" + using assms by (induct rule: equivclp_induct) (auto intro: ACI_map_respects equivclp_trans) + +lemma ACI_set_eq: + fixes r s :: "'a rexp" + assumes "r ~ s" + shows "set_rexp r = set_rexp s" + using assms by (induct r s rule: ACI.induct) auto + +lemma ACIcl_set_eq: + fixes r s :: "'a rexp" + assumes "r ~~ s" + shows "set_rexp r = set_rexp s" + using assms by (induct rule: equivclp_induct) (auto dest: ACI_set_eq) + +lemma Alt_eq_map_rexp_iff: + "Alt r s = map_rexp f x \ (\r' s'. x = Alt r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" + "map_rexp f x = Alt r s \ (\r' s'. x = Alt r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" + by (cases x; auto)+ + +lemma Conc_eq_map_rexp_iff: + "Conc r s = map_rexp f x \ (\r' s'. x = Conc r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" + "map_rexp f x = Conc r s \ (\r' s'. x = Conc r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" + by (cases x; auto)+ + +lemma Star_eq_map_rexp_iff: + "Star r = map_rexp f x \ (\r'. x = Star r' \ map_rexp f r' = r)" + "map_rexp f x = Star r \ (\r'. x = Star r' \ map_rexp f r' = r)" + by (cases x; auto)+ + +lemma AA: "r ~~ r' \ s ~~ s' \ Alt r s ~~ Alt r' s'" +proof (induct rule: equivclp_induct) + case base + then show ?case + by (induct rule: equivclp_induct) (auto elim!: equivclp_trans) +next + case (step s t) + then show ?case + by (auto elim!: equivclp_trans) +qed + +lemma AAA: "(~)\<^sup>*\<^sup>* r r' \ (~)\<^sup>*\<^sup>* s s' \ (~)\<^sup>*\<^sup>* (Alt r s) (Alt r' s')" +proof (induct r r' rule: rtranclp.induct) + case (rtrancl_refl r) + then show ?case + by (induct s s' rule: rtranclp.induct) + (auto elim!: rtranclp.rtrancl_into_rtrancl) +next + case (rtrancl_into_rtrancl a b c) + then show ?case + by (auto elim!: rtranclp.rtrancl_into_rtrancl) +qed + +lemma CC: "r ~~ r' \ s ~~ s' \ Conc r s ~~ Conc r' s'" +proof (induct rule: equivclp_induct) + case base + then show ?case + by (induct rule: equivclp_induct) (auto elim!: equivclp_trans) +next + case (step s t) + then show ?case + by (auto elim!: equivclp_trans) +qed + +lemma CCC: "(~)\<^sup>*\<^sup>* r r' \ (~)\<^sup>*\<^sup>* s s' \ (~)\<^sup>*\<^sup>* (Conc r s) (Conc r' s')" +proof (induct r r' rule: rtranclp.induct) + case (rtrancl_refl r) + then show ?case + by (induct s s' rule: rtranclp.induct) + (auto elim!: rtranclp.rtrancl_into_rtrancl) +next + case (rtrancl_into_rtrancl a b c) + then show ?case + by (auto elim!: rtranclp.rtrancl_into_rtrancl) +qed + +lemma SS: "r ~~ r' \ Star r ~~ Star r'" +proof (induct rule: equivclp_induct) + case (step s t) + then show ?case + by (auto elim!: equivclp_trans) +qed auto + +lemma SSS: "(~)\<^sup>*\<^sup>* r r' \ (~)\<^sup>*\<^sup>* (Star r) (Star r')" +proof (induct r r' rule: rtranclp.induct) + case (rtrancl_into_rtrancl a b c) + then show ?case + by (auto elim!: rtranclp.rtrancl_into_rtrancl) +qed auto + +lemma map_rexp_ACI_inv: "map_rexp f x ~ y \ \z. x ~~ z \ y = map_rexp f z" +proof (induct "map_rexp f x" y arbitrary: x rule: ACI.induct) + case (a1 r s t) + then obtain r' s' t' where "x = Alt (Alt r' s') t'" + "map_rexp f r' = r" "map_rexp f s' = s" "map_rexp f t' = t" + by (auto simp: Alt_eq_map_rexp_iff) + then show ?case + by (intro exI[of _ "Alt r' (Alt s' t')"]) auto +next + case (a2 r s t) + then obtain r' s' t' where "x = Alt r' (Alt s' t')" + "map_rexp f r' = r" "map_rexp f s' = s" "map_rexp f t' = t" + by (auto simp: Alt_eq_map_rexp_iff) + then show ?case + by (intro exI[of _ "Alt (Alt r' s') t'"]) auto +next + case (c r s) + then obtain r' s' where "x = Alt r' s'" + "map_rexp f r' = r" "map_rexp f s' = s" + by (auto simp: Alt_eq_map_rexp_iff) + then show ?case + by (intro exI[of _ "Alt s' r'"]) auto +next + case (i r) + then show ?case + by (intro exI[of _ "Alt r r"]) auto +next + case (R r) + then show ?case by (auto intro: exI[of _ r]) +next + case (A r rr s ss) + then obtain r' s' where "x = Alt r' s'" + "map_rexp f r' = r" "map_rexp f s' = s" + by (auto simp: Alt_eq_map_rexp_iff) + moreover from A(2)[OF this(2)[symmetric]] A(4)[OF this(3)[symmetric]] obtain rr' ss' where + "r' ~~ rr'" "rr = map_rexp f rr'" "s' ~~ ss'" "ss = map_rexp f ss'" + by blast + ultimately show ?case using A(1,3) + by (intro exI[of _ "Alt rr' ss'"]) (auto intro!: AA) +next + case (C r rr s ss) + then obtain r' s' where "x = Conc r' s'" + "map_rexp f r' = r" "map_rexp f s' = s" + by (auto simp: Conc_eq_map_rexp_iff) + moreover from C(2)[OF this(2)[symmetric]] C(4)[OF this(3)[symmetric]] obtain rr' ss' where + "r' ~~ rr'" "rr = map_rexp f rr'" "s' ~~ ss'" "ss = map_rexp f ss'" + by blast + ultimately show ?case using C(1,3) + by (intro exI[of _ "Conc rr' ss'"]) (auto intro!: CC) +next + case (S r rr) + then obtain r' where "x = Star r'" "map_rexp f r' = r" + by (auto simp: Star_eq_map_rexp_iff) + moreover from S(2)[OF this(2)[symmetric]] obtain rr' where "r' ~~ rr'" "rr = map_rexp f rr'" + by blast + ultimately show ?case + by (intro exI[of _ "Star rr'"]) (auto intro!: SS) +qed + +lemma reflclp_ACI: "(~)\<^sup>=\<^sup>= = (~)" + unfolding fun_eq_iff + by auto + +lemma strong_confluentp_ACI: "strong_confluentp (~)" + apply (rule strong_confluentpI, unfold reflclp_ACI) + subgoal for x y z + proof (induct x y arbitrary: z rule: ACI.induct) + case (a1 r s t) + then show ?case + by (auto intro: AAA converse_rtranclp_into_rtranclp) + next + case (a2 r s t) + then show ?case + by (auto intro: AAA converse_rtranclp_into_rtranclp) + next + case (c r s) + then show ?case + by (cases rule: ACI.cases) (auto intro: AAA) + next + case (i r) + then show ?case + by (auto intro: AAA) + next + case (R r) + then show ?case + by auto + next + note A1 = ACI.A[OF _ R] + note A2 = ACI.A[OF R] + case (A r r' s s') + from A(5,1-4) show ?case + proof (cases rule: ACI.cases) + case inner: (a1 r'' s'') + from A(1,3) show ?thesis + unfolding inner + proof (elim ACI.cases[of _ r'], goal_cases Xa1 Xa2 XC XI XR XP XT XS) + case (Xa1 rr ss tt) + with A(3) show ?case + by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A2[OF A2]], rotated]) + (metis a1 a2 A1 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) + next + case (Xa2 r s t) + then show ?case + by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A2[OF A2]], rotated]) + (metis a1 A1 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) + next + case (XC r s) + then show ?case + by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A2[OF A2]], rotated]) + (metis a1 c A1 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) + next + case (XI r) + then show ?case + apply (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ ACI.A[OF i ACI.A[OF i]]], rotated]) + apply hypsubst + apply (rule converse_rtranclp_into_rtranclp, rule a1) + apply (rule converse_rtranclp_into_rtranclp, rule a1) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule A1, rule c) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a1) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a1) + apply (rule converse_rtranclp_into_rtranclp, rule a2) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) + apply blast + done + qed auto + next + case inner: (a2 s'' t'') + with A(1,3) show ?thesis + unfolding inner + proof (elim ACI.cases[of _ s'], goal_cases Xa1 Xa2 XC XI XR XP XT XS) + case (Xa1 rr ss tt) + with A(3) show ?case + by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A1[OF A1]], rotated]) + (metis a2 A2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) + next + case (Xa2 r s t) + then show ?case + by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A1[OF A1]], rotated]) + (metis a1 a2 A2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) + next + case (XC r s) + then show ?case + by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A1[OF A1]], rotated]) + (metis a2 c A2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) + next + case (XI r) + then show ?case + apply (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ ACI.A[OF ACI.A[OF _ i] i]], rotated]) + apply hypsubst + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule A1, rule c) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a1) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule A2, rule a1) + apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) + apply (rule converse_rtranclp_into_rtranclp, rule a2) + apply blast + done + qed auto + qed (auto 5 0 intro: AAA) + next + case (C r r' s s') + from C(5,1-4) show ?case + by (cases rule: ACI.cases) (auto 5 0 intro: CCC) + next + case (S r r') + from S(3,1,2) show ?case + by (cases rule: ACI.cases) (auto intro: SSS) + qed + done + +lemma confluent_quotient_ACI: + "confluent_quotient (~) (~~) (~~) (~~) (~~) (~~) + (map_rexp fst) (map_rexp snd) (map_rexp fst) (map_rexp snd) + rel_rexp rel_rexp rel_rexp set_rexp set_rexp" + by unfold_locales (auto dest: ACIcl_set_eq simp: rexp.in_rel rexp.rel_compp map_rexp_ACI_inv + intro: equivpI reflpI sympI transpI ACIcl_map_respects + strong_confluentp_imp_confluentp[OF strong_confluentp_ACI]) + +inductive ACIEQ (infix "\" 64) where + a: "Alt (Alt r s) t \ Alt r (Alt s t)" +| c: "Alt r s \ Alt s r" +| i: "Alt r r \ r" +| A: "r \ r' \ s \ s' \ Alt r s \ Alt r' s'" +| C: "r \ r' \ s \ s' \ Conc r s \ Conc r' s'" +| S: "r \ r' \ Star r \ Star r'" +| r: "r \ r" +| s: "r \ s \ s \ r" +| t: "r \ s \ s \ t \ r \ t" + +lemma ACIEQ_le_ACIcl: "r \ s \ r ~~ s" + by (induct r s rule: ACIEQ.induct) (auto intro: AA CC SS equivclp_sym equivclp_trans) + +lemma ACI_le_ACIEQ: "r ~ s \ r \ s" + by (induct r s rule: ACI.induct) (auto intro: ACIEQ.intros) + +lemma ACIcl_le_ACIEQ: "r ~~ s \ r \ s" + by (induct rule: equivclp_induct) (auto 0 3 intro: ACIEQ.intros ACI_le_ACIEQ) + +lemma ACIEQ_alt: "(\) = (~~)" + by (simp add: ACIEQ_le_ACIcl ACIcl_le_ACIEQ antisym predicate2I) + +quotient_type 'a ACI_rexp = "'a rexp" / "(\)" + unfolding ACIEQ_alt by (auto intro!: equivpI reflpI sympI transpI) + +lift_bnf (no_warn_wits) 'a ACI_rexp + unfolding ACIEQ_alt + subgoal for A Q by (rule confluent_quotient.subdistributivity[OF confluent_quotient_ACI]) + subgoal for Ss by (intro eq_set_preserves_inter[rotated] ACIcl_set_eq; auto) + done + +datatype ldl = Prop string | And ldl ldl | Not ldl | Match "ldl ACI_rexp" + +end \ No newline at end of file diff --git a/src/HOL/Equiv_Relations.thy b/src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy +++ b/src/HOL/Equiv_Relations.thy @@ -1,489 +1,588 @@ (* Title: HOL/Equiv_Relations.thy Author: Lawrence C Paulson, 1996 Cambridge University Computer Laboratory *) section \Equivalence Relations in Higher-Order Set Theory\ theory Equiv_Relations imports Groups_Big begin subsection \Equivalence relations -- set version\ definition equiv :: "'a set \ ('a \ 'a) set \ bool" where "equiv A r \ refl_on A r \ sym r \ trans r" lemma equivI: "refl_on A r \ sym r \ trans r \ equiv A r" by (simp add: equiv_def) lemma equivE: assumes "equiv A r" obtains "refl_on A r" and "sym r" and "trans r" using assms by (simp add: equiv_def) text \ Suppes, Theorem 70: \r\ is an equiv relation iff \r\ O r = r\. First half: \equiv A r \ r\ O r = r\. \ lemma sym_trans_comp_subset: "sym r \ trans r \ r\ O r \ r" unfolding trans_def sym_def converse_unfold by blast lemma refl_on_comp_subset: "refl_on A r \ r \ r\ O r" unfolding refl_on_def by blast lemma equiv_comp_eq: "equiv A r \ r\ O r = r" apply (unfold equiv_def) apply clarify apply (rule equalityI) apply (iprover intro: sym_trans_comp_subset refl_on_comp_subset)+ done text \Second half.\ lemma comp_equivI: "r\ O r = r \ Domain r = A \ equiv A r" apply (unfold equiv_def refl_on_def sym_def trans_def) apply (erule equalityE) apply (subgoal_tac "\x y. (x, y) \ r \ (y, x) \ r") apply fast apply fast done subsection \Equivalence classes\ lemma equiv_class_subset: "equiv A r \ (a, b) \ r \ r``{a} \ r``{b}" \ \lemma for the next result\ unfolding equiv_def trans_def sym_def by blast theorem equiv_class_eq: "equiv A r \ (a, b) \ r \ r``{a} = r``{b}" apply (assumption | rule equalityI equiv_class_subset)+ apply (unfold equiv_def sym_def) apply blast done lemma equiv_class_self: "equiv A r \ a \ A \ a \ r``{a}" unfolding equiv_def refl_on_def by blast lemma subset_equiv_class: "equiv A r \ r``{b} \ r``{a} \ b \ A \ (a, b) \ r" \ \lemma for the next result\ unfolding equiv_def refl_on_def by blast lemma eq_equiv_class: "r``{a} = r``{b} \ equiv A r \ b \ A \ (a, b) \ r" by (iprover intro: equalityD2 subset_equiv_class) lemma equiv_class_nondisjoint: "equiv A r \ x \ (r``{a} \ r``{b}) \ (a, b) \ r" unfolding equiv_def trans_def sym_def by blast lemma equiv_type: "equiv A r \ r \ A \ A" unfolding equiv_def refl_on_def by blast lemma equiv_class_eq_iff: "equiv A r \ (x, y) \ r \ r``{x} = r``{y} \ x \ A \ y \ A" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) lemma eq_equiv_class_iff: "equiv A r \ x \ A \ y \ A \ r``{x} = r``{y} \ (x, y) \ r" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) subsection \Quotients\ definition quotient :: "'a set \ ('a \ 'a) set \ 'a set set" (infixl "'/'/" 90) where "A//r = (\x \ A. {r``{x}})" \ \set of equiv classes\ lemma quotientI: "x \ A ==> r``{x} \ A//r" unfolding quotient_def by blast lemma quotientE: "X \ A//r \ (\x. X = r``{x} \ x \ A \ P) \ P" unfolding quotient_def by blast lemma Union_quotient: "equiv A r \ \(A//r) = A" unfolding equiv_def refl_on_def quotient_def by blast lemma quotient_disj: "equiv A r \ X \ A//r \ Y \ A//r \ X = Y \ X \ Y = {}" apply (unfold quotient_def) apply clarify apply (rule equiv_class_eq) apply assumption apply (unfold equiv_def trans_def sym_def) apply blast done lemma quotient_eqI: "equiv A r \ X \ A//r \ Y \ A//r \ x \ X \ y \ Y \ (x, y) \ r \ X = Y" apply (clarify elim!: quotientE) apply (rule equiv_class_eq) apply assumption apply (unfold equiv_def sym_def trans_def) apply blast done lemma quotient_eq_iff: "equiv A r \ X \ A//r \ Y \ A//r \ x \ X \ y \ Y \ X = Y \ (x, y) \ r" apply (rule iffI) prefer 2 apply (blast del: equalityI intro: quotient_eqI) apply (clarify elim!: quotientE) apply (unfold equiv_def sym_def trans_def) apply blast done lemma eq_equiv_class_iff2: "equiv A r \ x \ A \ y \ A \ {x}//r = {y}//r \ (x, y) \ r" by (simp add: quotient_def eq_equiv_class_iff) lemma quotient_empty [simp]: "{}//r = {}" by (simp add: quotient_def) lemma quotient_is_empty [iff]: "A//r = {} \ A = {}" by (simp add: quotient_def) lemma quotient_is_empty2 [iff]: "{} = A//r \ A = {}" by (simp add: quotient_def) lemma singleton_quotient: "{x}//r = {r `` {x}}" by (simp add: quotient_def) lemma quotient_diff1: "inj_on (\a. {a}//r) A \ a \ A \ (A - {a})//r = A//r - {a}//r" unfolding quotient_def inj_on_def by blast subsection \Refinement of one equivalence relation WRT another\ lemma refines_equiv_class_eq: "R \ S \ equiv A R \ equiv A S \ R``(S``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_class_eq2: "R \ S \ equiv A R \ equiv A S \ S``(R``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_image_eq: "R \ S \ equiv A R \ equiv A S \ (\X. S``X) ` (A//R) = A//S" by (auto simp: quotient_def image_UN refines_equiv_class_eq2) lemma finite_refines_finite: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ finite (A//S)" by (erule finite_surj [where f = "\X. S``X"]) (simp add: refines_equiv_image_eq) lemma finite_refines_card_le: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ card (A//S) \ card (A//R)" by (subst refines_equiv_image_eq [of R S A, symmetric]) (auto simp: card_image_le [where f = "\X. S``X"]) subsection \Defining unary operations upon equivalence classes\ text \A congruence-preserving function.\ definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where "congruent r f \ (\(y, z) \ r. f y = f z)" lemma congruentI: "(\y z. (y, z) \ r \ f y = f z) \ congruent r f" by (auto simp add: congruent_def) lemma congruentD: "congruent r f \ (y, z) \ r \ f y = f z" by (auto simp add: congruent_def) abbreviation RESPECTS :: "('a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects" 80) where "f respects r \ congruent r f" lemma UN_constant_eq: "a \ A \ \y \ A. f y = c \ (\y \ A. f y) = c" \ \lemma required to prove \UN_equiv_class\\ by auto lemma UN_equiv_class: "equiv A r \ f respects r \ a \ A \ (\x \ r``{a}. f x) = f a" \ \Conversion rule\ apply (rule equiv_class_self [THEN UN_constant_eq]) apply assumption apply assumption apply (unfold equiv_def congruent_def sym_def) apply (blast del: equalityI) done lemma UN_equiv_class_type: "equiv A r \ f respects r \ X \ A//r \ (\x. x \ A \ f x \ B) \ (\x \ X. f x) \ B" apply (unfold quotient_def) apply clarify apply (subst UN_equiv_class) apply auto done text \ Sufficient conditions for injectiveness. Could weaken premises! major premise could be an inclusion; \bcong\ could be \\y. y \ A \ f y \ B\. \ lemma UN_equiv_class_inject: "equiv A r \ f respects r \ (\x \ X. f x) = (\y \ Y. f y) \ X \ A//r ==> Y \ A//r \ (\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r) \ X = Y" apply (unfold quotient_def) apply clarify apply (rule equiv_class_eq) apply assumption apply (subgoal_tac "f x = f xa") apply blast apply (erule box_equals) apply (assumption | rule UN_equiv_class)+ done subsection \Defining binary operations upon equivalence classes\ text \A congruence-preserving function of two arguments.\ definition congruent2 :: "('a \ 'a) set \ ('b \ 'b) set \ ('a \ 'b \ 'c) \ bool" where "congruent2 r1 r2 f \ (\(y1, z1) \ r1. \(y2, z2) \ r2. f y1 y2 = f z1 z2)" lemma congruent2I': assumes "\y1 z1 y2 z2. (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" shows "congruent2 r1 r2 f" using assms by (auto simp add: congruent2_def) lemma congruent2D: "congruent2 r1 r2 f \ (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" by (auto simp add: congruent2_def) text \Abbreviation for the common case where the relations are identical.\ abbreviation RESPECTS2:: "('a \ 'a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects2" 80) where "f respects2 r \ congruent2 r r f" lemma congruent2_implies_congruent: "equiv A r1 \ congruent2 r1 r2 f \ a \ A \ congruent r2 (f a)" unfolding congruent_def congruent2_def equiv_def refl_on_def by blast lemma congruent2_implies_congruent_UN: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a \ A2 \ congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" apply (unfold congruent_def) apply clarify apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+) apply (simp add: UN_equiv_class congruent2_implies_congruent) apply (unfold congruent2_def equiv_def refl_on_def) apply (blast del: equalityI) done lemma UN_equiv_class2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a1 \ A1 \ a2 \ A2 \ (\x1 \ r1``{a1}. \x2 \ r2``{a2}. f x1 x2) = f a1 a2" by (simp add: UN_equiv_class congruent2_implies_congruent congruent2_implies_congruent_UN) lemma UN_equiv_class_type2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ X1 \ A1//r1 \ X2 \ A2//r2 \ (\x1 x2. x1 \ A1 \ x2 \ A2 \ f x1 x2 \ B) \ (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" apply (unfold quotient_def) apply clarify apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN congruent2_implies_congruent quotientI) done lemma UN_UN_split_split_eq: "(\(x1, x2) \ X. \(y1, y2) \ Y. A x1 x2 y1 y2) = (\x \ X. \y \ Y. (\(x1, x2). (\(y1, y2). A x1 x2 y1 y2) y) x)" \ \Allows a natural expression of binary operators,\ \ \without explicit calls to \split\\ by auto lemma congruent2I: "equiv A1 r1 \ equiv A2 r2 \ (\y z w. w \ A2 \ (y,z) \ r1 \ f y w = f z w) \ (\y z w. w \ A1 \ (y,z) \ r2 \ f w y = f w z) \ congruent2 r1 r2 f" \ \Suggested by John Harrison -- the two subproofs may be\ \ \\<^emph>\much\ simpler than the direct proof.\ apply (unfold congruent2_def equiv_def refl_on_def) apply clarify apply (blast intro: trans) done lemma congruent2_commuteI: assumes equivA: "equiv A r" and commute: "\y z. y \ A \ z \ A \ f y z = f z y" and congt: "\y z w. w \ A \ (y,z) \ r \ f w y = f w z" shows "f respects2 r" apply (rule congruent2I [OF equivA equivA]) apply (rule commute [THEN trans]) apply (rule_tac [3] commute [THEN trans, symmetric]) apply (rule_tac [5] sym) apply (rule congt | assumption | erule equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2])+ done subsection \Quotients and finiteness\ text \Suggested by Florian Kammüller\ lemma finite_quotient: "finite A \ r \ A \ A \ finite (A//r)" \ \recall @{thm equiv_type}\ apply (rule finite_subset) apply (erule_tac [2] finite_Pow_iff [THEN iffD2]) apply (unfold quotient_def) apply blast done lemma finite_equiv_class: "finite A \ r \ A \ A \ X \ A//r \ finite X" apply (unfold quotient_def) apply (rule finite_subset) prefer 2 apply assumption apply blast done lemma equiv_imp_dvd_card: "finite A \ equiv A r \ \X \ A//r. k dvd card X \ k dvd card A" apply (rule Union_quotient [THEN subst [where P="\A. k dvd card A"]]) apply assumption apply (rule dvd_partition) prefer 3 apply (blast dest: quotient_disj) apply (simp_all add: Union_quotient equiv_type) done lemma card_quotient_disjoint: "finite A \ inj_on (\x. {x} // r) A \ card (A//r) = card A" apply (simp add:quotient_def) apply (subst card_UN_disjoint) apply assumption apply simp apply (fastforce simp add:inj_on_def) apply simp done subsection \Projection\ definition proj :: "('b \ 'a) set \ 'b \ 'a set" where "proj r x = r `` {x}" lemma proj_preserves: "x \ A \ proj r x \ A//r" unfolding proj_def by (rule quotientI) lemma proj_in_iff: assumes "equiv A r" shows "proj r x \ A//r \ x \ A" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: proj_preserves) next assume ?lhs then show ?rhs unfolding proj_def quotient_def proof clarsimp fix y assume y: "y \ A" and "r `` {x} = r `` {y}" moreover have "y \ r `` {y}" using assms y unfolding equiv_def refl_on_def by blast ultimately have "(x, y) \ r" by blast then show "x \ A" using assms unfolding equiv_def refl_on_def by blast qed qed lemma proj_iff: "equiv A r \ {x, y} \ A \ proj r x = proj r y \ (x, y) \ r" by (simp add: proj_def eq_equiv_class_iff) (* lemma in_proj: "\equiv A r; x \ A\ \ x \ proj r x" unfolding proj_def equiv_def refl_on_def by blast *) lemma proj_image: "proj r ` A = A//r" unfolding proj_def[abs_def] quotient_def by blast lemma in_quotient_imp_non_empty: "equiv A r \ X \ A//r \ X \ {}" unfolding quotient_def using equiv_class_self by fast lemma in_quotient_imp_in_rel: "equiv A r \ X \ A//r \ {x, y} \ X \ (x, y) \ r" using quotient_eq_iff[THEN iffD1] by fastforce lemma in_quotient_imp_closed: "equiv A r \ X \ A//r \ x \ X \ (x, y) \ r \ y \ X" unfolding quotient_def equiv_def trans_def by blast lemma in_quotient_imp_subset: "equiv A r \ X \ A//r \ X \ A" using in_quotient_imp_in_rel equiv_type by fastforce subsection \Equivalence relations -- predicate version\ text \Partial equivalences.\ definition part_equivp :: "('a \ 'a \ bool) \ bool" where "part_equivp R \ (\x. R x x) \ (\x y. R x y \ R x x \ R y y \ R x = R y)" \ \John-Harrison-style characterization\ lemma part_equivpI: "\x. R x x \ symp R \ transp R \ part_equivp R" by (auto simp add: part_equivp_def) (auto elim: sympE transpE) lemma part_equivpE: assumes "part_equivp R" obtains x where "R x x" and "symp R" and "transp R" proof - from assms have 1: "\x. R x x" and 2: "\x y. R x y \ R x x \ R y y \ R x = R y" unfolding part_equivp_def by blast+ from 1 obtain x where "R x x" .. moreover have "symp R" proof (rule sympI) fix x y assume "R x y" with 2 [of x y] show "R y x" by auto qed moreover have "transp R" proof (rule transpI) fix x y z assume "R x y" and "R y z" with 2 [of x y] 2 [of y z] show "R x z" by auto qed ultimately show thesis by (rule that) qed lemma part_equivp_refl_symp_transp: "part_equivp R \ (\x. R x x) \ symp R \ transp R" by (auto intro: part_equivpI elim: part_equivpE) lemma part_equivp_symp: "part_equivp R \ R x y \ R y x" by (erule part_equivpE, erule sympE) lemma part_equivp_transp: "part_equivp R \ R x y \ R y z \ R x z" by (erule part_equivpE, erule transpE) lemma part_equivp_typedef: "part_equivp R \ \d. d \ {c. \x. R x x \ c = Collect (R x)}" by (auto elim: part_equivpE) text \Total equivalences.\ definition equivp :: "('a \ 'a \ bool) \ bool" where "equivp R \ (\x y. R x y = (R x = R y))" \ \John-Harrison-style characterization\ lemma equivpI: "reflp R \ symp R \ transp R \ equivp R" by (auto elim: reflpE sympE transpE simp add: equivp_def) lemma equivpE: assumes "equivp R" obtains "reflp R" and "symp R" and "transp R" using assms by (auto intro!: that reflpI sympI transpI simp add: equivp_def) lemma equivp_implies_part_equivp: "equivp R \ part_equivp R" by (auto intro: part_equivpI elim: equivpE reflpE) lemma equivp_equiv: "equiv UNIV A \ equivp (\x y. (x, y) \ A)" by (auto intro!: equivI equivpI [to_set] elim!: equivE equivpE [to_set]) lemma equivp_reflp_symp_transp: "equivp R \ reflp R \ symp R \ transp R" by (auto intro: equivpI elim: equivpE) lemma identity_equivp: "equivp (=)" by (auto intro: equivpI reflpI sympI transpI) lemma equivp_reflp: "equivp R \ R x x" by (erule equivpE, erule reflpE) lemma equivp_symp: "equivp R \ R x y \ R y x" by (erule equivpE, erule sympE) lemma equivp_transp: "equivp R \ R x y \ R y z \ R x z" by (erule equivpE, erule transpE) +lemma equivp_rtranclp: "symp r \ equivp r\<^sup>*\<^sup>*" + by(intro equivpI reflpI sympI transpI)(auto dest: sympD[OF symp_rtranclp]) + +lemmas equivp_rtranclp_symclp [simp] = equivp_rtranclp[OF symp_symclp] + +lemma equivp_vimage2p: "equivp R \ equivp (vimage2p f f R)" + by(auto simp add: equivp_def vimage2p_def dest: fun_cong) + +lemma equivp_imp_transp: "equivp R \ transp R" + by(simp add: equivp_reflp_symp_transp) + + +subsection \Equivalence closure\ + +definition equivclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" where + "equivclp r = (symclp r)\<^sup>*\<^sup>*" + +lemma transp_equivclp [simp]: "transp (equivclp r)" + by(simp add: equivclp_def) + +lemma reflp_equivclp [simp]: "reflp (equivclp r)" + by(simp add: equivclp_def) + +lemma symp_equivclp [simp]: "symp (equivclp r)" + by(simp add: equivclp_def) + +lemma equivp_evquivclp [simp]: "equivp (equivclp r)" + by(simp add: equivpI) + +lemma tranclp_equivclp [simp]: "(equivclp r)\<^sup>+\<^sup>+ = equivclp r" + by(simp add: equivclp_def) + +lemma rtranclp_equivclp [simp]: "(equivclp r)\<^sup>*\<^sup>* = equivclp r" + by(simp add: equivclp_def) + +lemma symclp_equivclp [simp]: "symclp (equivclp r) = equivclp r" + by(simp add: equivclp_def symp_symclp_eq) + +lemma equivclp_symclp [simp]: "equivclp (symclp r) = equivclp r" + by(simp add: equivclp_def) + +lemma equivclp_conversep [simp]: "equivclp (conversep r) = equivclp r" + by(simp add: equivclp_def) + +lemma equivclp_sym [sym]: "equivclp r x y \ equivclp r y x" + by(rule sympD[OF symp_equivclp]) + +lemma equivclp_OO_equivclp_le_equivclp: "equivclp r OO equivclp r \ equivclp r" + by(rule transp_relcompp_less_eq transp_equivclp)+ + +lemma rtranlcp_le_equivclp: "r\<^sup>*\<^sup>* \ equivclp r" + unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) + +lemma rtranclp_conversep_le_equivclp: "r\\\<^sup>*\<^sup>* \ equivclp r" + unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) + +lemma symclp_rtranclp_le_equivclp: "symclp r\<^sup>*\<^sup>* \ equivclp r" + unfolding symclp_pointfree + by(rule le_supI)(simp_all add: rtranclp_conversep[symmetric] rtranlcp_le_equivclp rtranclp_conversep_le_equivclp) + +lemma r_OO_conversep_into_equivclp: + "r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>* \ equivclp r" + by(blast intro: order_trans[OF _ equivclp_OO_equivclp_le_equivclp] relcompp_mono rtranlcp_le_equivclp rtranclp_conversep_le_equivclp del: predicate2I) + +lemma equivclp_induct [consumes 1, case_names base step, induct pred: equivclp]: + assumes a: "equivclp r a b" + and cases: "P a" "\y z. equivclp r a y \ r y z \ r z y \ P y \ P z" + shows "P b" + using a unfolding equivclp_def + by(induction rule: rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) + +lemma converse_equivclp_induct [consumes 1, case_names base step]: + assumes major: "equivclp r a b" + and cases: "P b" "\y z. r y z \ r z y \ equivclp r z b \ P z \ P y" + shows "P a" + using major unfolding equivclp_def + by(induction rule: converse_rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) + +lemma equivclp_refl [simp]: "equivclp r x x" + by(rule reflpD[OF reflp_equivclp]) + +lemma r_into_equivclp [intro]: "r x y \ equivclp r x y" + unfolding equivclp_def by(blast intro: symclpI) + +lemma converse_r_into_equivclp [intro]: "r y x \ equivclp r x y" + unfolding equivclp_def by(blast intro: symclpI) + +lemma rtranclp_into_equivclp: "r\<^sup>*\<^sup>* x y \ equivclp r x y" + using rtranlcp_le_equivclp[of r] by blast + +lemma converse_rtranclp_into_equivclp: "r\<^sup>*\<^sup>* y x \ equivclp r x y" + by(blast intro: equivclp_sym rtranclp_into_equivclp) + +lemma equivclp_into_equivclp: "\ equivclp r a b; r b c \ r c b \ \ equivclp r a c" + unfolding equivclp_def by(erule rtranclp.rtrancl_into_rtrancl)(auto intro: symclpI) + +lemma equivclp_trans [trans]: "\ equivclp r a b; equivclp r b c \ \ equivclp r a c" + using equivclp_OO_equivclp_le_equivclp[of r] by blast + hide_const (open) proj end diff --git a/src/HOL/Library/Confluence.thy b/src/HOL/Library/Confluence.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Library/Confluence.thy @@ -0,0 +1,77 @@ +theory Confluence imports + Main +begin + +section \Confluence\ + +definition semiconfluentp :: "('a \ 'a \ bool) \ bool" where + "semiconfluentp r \ r\\ OO r\<^sup>*\<^sup>* \ r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>*" + +definition confluentp :: "('a \ 'a \ bool) \ bool" where + "confluentp r \ r\\\<^sup>*\<^sup>* OO r\<^sup>*\<^sup>* \ r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>*" + +definition strong_confluentp :: "('a \ 'a \ bool) \ bool" where + "strong_confluentp r \ r\\ OO r \ r\<^sup>*\<^sup>* OO (r\\)\<^sup>=\<^sup>=" + +lemma semiconfluentpI [intro?]: + "semiconfluentp r" if "\x y z. \ r x y; r\<^sup>*\<^sup>* x z \ \ \u. r\<^sup>*\<^sup>* y u \ r\<^sup>*\<^sup>* z u" + using that unfolding semiconfluentp_def rtranclp_conversep by blast + +lemma semiconfluentpD: "\u. r\<^sup>*\<^sup>* y u \ r\<^sup>*\<^sup>* z u" if "semiconfluentp r" "r x y" "r\<^sup>*\<^sup>* x z" + using that unfolding semiconfluentp_def rtranclp_conversep by blast + +lemma confluentpI: + "confluentp r" if "\x y z. \ r\<^sup>*\<^sup>* x y; r\<^sup>*\<^sup>* x z \ \ \u. r\<^sup>*\<^sup>* y u \ r\<^sup>*\<^sup>* z u" + using that unfolding confluentp_def rtranclp_conversep by blast + +lemma confluentpD: "\u. r\<^sup>*\<^sup>* y u \ r\<^sup>*\<^sup>* z u" if "confluentp r" "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x z" + using that unfolding confluentp_def rtranclp_conversep by blast + +lemma strong_confluentpI [intro?]: + "strong_confluentp r" if "\x y z. \ r x y; r x z \ \ \u. r\<^sup>*\<^sup>* y u \ r\<^sup>=\<^sup>= z u" + using that unfolding strong_confluentp_def by blast + +lemma strong_confluentpD: "\u. r\<^sup>*\<^sup>* y u \ r\<^sup>=\<^sup>= z u" if "strong_confluentp r" "r x y" "r x z" + using that unfolding strong_confluentp_def by blast + +lemma semiconfluentp_imp_confluentp: "confluentp r" if r: "semiconfluentp r" +proof(rule confluentpI) + show "\u. r\<^sup>*\<^sup>* y u \ r\<^sup>*\<^sup>* z u" if "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x z" for x y z + using that(2,1) + by(induction arbitrary: y rule: converse_rtranclp_induct) + (blast intro: rtranclp_trans dest: r[THEN semiconfluentpD])+ +qed + +lemma confluentp_imp_semiconfluentp: "semiconfluentp r" if "confluentp r" + using that by(auto intro!: semiconfluentpI dest: confluentpD[OF that]) + +lemma confluentp_eq_semiconfluentp: "confluentp r \ semiconfluentp r" + by(blast intro: semiconfluentp_imp_confluentp confluentp_imp_semiconfluentp) + +lemma confluentp_conv_strong_confluentp_rtranclp: + "confluentp r \ strong_confluentp (r\<^sup>*\<^sup>*)" + by(auto simp add: confluentp_def strong_confluentp_def rtranclp_conversep) + +lemma strong_confluentp_into_semiconfluentp: + "semiconfluentp r" if r: "strong_confluentp r" +proof + show "\u. r\<^sup>*\<^sup>* y u \ r\<^sup>*\<^sup>* z u" if "r x y" "r\<^sup>*\<^sup>* x z" for x y z + using that(2,1) + apply(induction arbitrary: y rule: converse_rtranclp_induct) + subgoal by blast + subgoal for a b c + by (drule (1) strong_confluentpD[OF r, of a c])(auto 10 0 intro: rtranclp_trans) + done +qed + +lemma strong_confluentp_imp_confluentp: "confluentp r" if "strong_confluentp r" + unfolding confluentp_eq_semiconfluentp using that by(rule strong_confluentp_into_semiconfluentp) + +lemma semiconfluentp_equivclp: "equivclp r = r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>*" if r: "semiconfluentp r" +proof(rule antisym[rotated] r_OO_conversep_into_equivclp predicate2I)+ + show "(r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>*) x y" if "equivclp r x y" for x y using that unfolding equivclp_def rtranclp_conversep + by(induction rule: converse_rtranclp_induct) + (blast elim!: symclpE intro: converse_rtranclp_into_rtranclp rtranclp_trans dest: semiconfluentpD[OF r])+ +qed + +end \ No newline at end of file diff --git a/src/HOL/Library/Confluent_Quotient.thy b/src/HOL/Library/Confluent_Quotient.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Library/Confluent_Quotient.thy @@ -0,0 +1,74 @@ +theory Confluent_Quotient imports + Confluence +begin + +section \Subdistributivity for quotients via confluence\ + +locale confluent_quotient = + fixes R :: "'Fb \ 'Fb \ bool" + and Ea :: "'Fa \ 'Fa \ bool" + and Eb :: "'Fb \ 'Fb \ bool" + and Ec :: "'Fc \ 'Fc \ bool" + and Eab :: "'Fab \ 'Fab \ bool" + and Ebc :: "'Fbc \ 'Fbc \ bool" + and \_Faba :: "'Fab \ 'Fa" + and \_Fabb :: "'Fab \ 'Fb" + and \_Fbcb :: "'Fbc \ 'Fb" + and \_Fbcc :: "'Fbc \ 'Fc" + and rel_Fab :: "('a \ 'b \ bool) \ 'Fa \ 'Fb \ bool" + and rel_Fbc :: "('b \ 'c \ bool) \ 'Fb \ 'Fc \ bool" + and rel_Fac :: "('a \ 'c \ bool) \ 'Fa \ 'Fc \ bool" + and set_Fab :: "'Fab \ ('a \ 'b) set" + and set_Fbc :: "'Fbc \ ('b \ 'c) set" + assumes confluent: "confluentp R" + and retract1_ab: "\x y. R (\_Fabb x) y \ \z. Eab x z \ y = \_Fabb z" + and retract1_bc: "\x y. R (\_Fbcb x) y \ \z. Ebc x z \ y = \_Fbcb z" + and generated: "Eb \ equivclp R" + and set_ab: "\x y. Eab x y \ set_Fab x = set_Fab y" + and set_bc: "\x y. Ebc x y \ set_Fbc x = set_Fbc y" + and transp_a: "transp Ea" + and transp_c: "transp Ec" + and equivp_ab: "equivp Eab" + and equivp_bc: "equivp Ebc" + and in_rel_Fab: "\A x y. rel_Fab A x y \ (\z. z \ {x. set_Fab x \ {(x, y). A x y}} \ \_Faba z = x \ \_Fabb z = y)" + and in_rel_Fbc: "\B x y. rel_Fbc B x y \ (\z. z \ {x. set_Fbc x \ {(x, y). B x y}} \ \_Fbcb z = x \ \_Fbcc z = y)" + and rel_compp: "\A B. rel_Fac (A OO B) = rel_Fab A OO rel_Fbc B" + and \_Faba_respect: "rel_fun Eab Ea \_Faba \_Faba" + and \_Fbcc_respect: "rel_fun Ebc Ec \_Fbcc \_Fbcc" +begin + +lemma retract_ab: "R\<^sup>*\<^sup>* (\_Fabb x) y \ \z. Eab x z \ y = \_Fabb z" + by(induction rule: rtranclp_induct)(blast dest: retract1_ab intro: equivp_transp[OF equivp_ab] equivp_reflp[OF equivp_ab])+ + +lemma retract_bc: "R\<^sup>*\<^sup>* (\_Fbcb x) y \ \z. Ebc x z \ y = \_Fbcb z" + by(induction rule: rtranclp_induct)(blast dest: retract1_bc intro: equivp_transp[OF equivp_bc] equivp_reflp[OF equivp_bc])+ + +lemma subdistributivity: "rel_Fab A OO Eb OO rel_Fbc B \ Ea OO rel_Fac (A OO B) OO Ec" +proof(rule predicate2I; elim relcomppE) + fix x y y' z + assume "rel_Fab A x y" and "Eb y y'" and "rel_Fbc B y' z" + then obtain xy y'z + where xy: "set_Fab xy \ {(a, b). A a b}" "x = \_Faba xy" "y = \_Fabb xy" + and y'z: "set_Fbc y'z \ {(a, b). B a b}" "y' = \_Fbcb y'z" "z = \_Fbcc y'z" + by(auto simp add: in_rel_Fab in_rel_Fbc) + from \Eb y y'\ have "equivclp R y y'" using generated by blast + then obtain u where u: "R\<^sup>*\<^sup>* y u" "R\<^sup>*\<^sup>* y' u" + unfolding semiconfluentp_equivclp[OF confluent[THEN confluentp_imp_semiconfluentp]] + by(auto simp add: rtranclp_conversep) + with xy y'z obtain xy' y'z' + where retract1: "Eab xy xy'" "\_Fabb xy' = u" + and retract2: "Ebc y'z y'z'" "\_Fbcb y'z' = u" + by(auto dest!: retract_ab retract_bc) + from retract1(1) xy have "Ea x (\_Faba xy')" by(auto dest: \_Faba_respect[THEN rel_funD]) + moreover have "rel_Fab A (\_Faba xy') u" using xy retract1 + by(auto simp add: in_rel_Fab dest: set_ab) + moreover have "rel_Fbc B u (\_Fbcc y'z')" using y'z retract2 + by(auto simp add: in_rel_Fbc dest: set_bc) + moreover have "Ec (\_Fbcc y'z') z" using retract2 y'z equivp_symp[OF equivp_bc] + by(auto dest: \_Fbcc_respect[THEN rel_funD]) + ultimately show "(Ea OO rel_Fac (A OO B) OO Ec) x z" unfolding rel_compp by blast +qed + +end + +end \ No newline at end of file diff --git a/src/HOL/Library/Library.thy b/src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy +++ b/src/HOL/Library/Library.thy @@ -1,101 +1,103 @@ (*<*) theory Library imports AList Adhoc_Overloading BigO BNF_Axiomatization BNF_Corec Boolean_Algebra Bourbaki_Witt_Fixpoint Char_ord Code_Lazy Code_Test Combine_PER Complete_Partial_Order2 Conditional_Parametricity + Confluence + Confluent_Quotient Countable Countable_Complete_Lattices Countable_Set_Type Debug Diagonal_Subsequence Discrete Disjoint_Sets Dlist Dual_Ordered_Lattice Equipollence Extended Extended_Nat Extended_Nonnegative_Real Extended_Real Finite_Map Float FSet FuncSet Function_Division Fun_Lexorder Going_To_Filter Groups_Big_Fun Indicator_Function Infinite_Set Interval Interval_Float IArray Landau_Symbols Lattice_Algebras Lattice_Syntax Lattice_Constructions Linear_Temporal_Logic_on_Streams ListVector Lub_Glb Mapping Monad_Syntax More_List Multiset_Order Multiset_Permutations Nonpos_Ints Numeral_Type Omega_Words_Fun Open_State_Syntax Option_ord Order_Continuity Parallel Pattern_Aliases Periodic_Fun Perm Permutation Permutations Poly_Mapping Power_By_Squaring Preorder Product_Plus Quadratic_Discriminant Quotient_List Quotient_Option Quotient_Product Quotient_Set Quotient_Sum Quotient_Syntax Quotient_Type Ramsey Reflection Rewrite Saturated Set_Algebras Set_Idioms State_Monad Stirling Stream Sorting_Algorithms Sublist Sum_of_Squares Transitive_Closure_Table Tree_Multiset Tree_Real Type_Length Uprod While_Combinator Z2 begin end (*>*) diff --git a/src/HOL/List.thy b/src/HOL/List.thy --- a/src/HOL/List.thy +++ b/src/HOL/List.thy @@ -1,7521 +1,7566 @@ (* Title: HOL/List.thy Author: Tobias Nipkow; proofs tidied by LCP *) section \The datatype of finite lists\ theory List imports Sledgehammer Code_Numeral Lifting_Set begin datatype (set: 'a) list = Nil ("[]") | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) for map: map rel: list_all2 pred: list_all where "tl [] = []" datatype_compat list lemma [case_names Nil Cons, cases type: list]: \ \for backward compatibility -- names of variables differ\ "(y = [] \ P) \ (\a list. y = a # list \ P) \ P" by (rule list.exhaust) lemma [case_names Nil Cons, induct type: list]: \ \for backward compatibility -- names of variables differ\ "P [] \ (\a list. P list \ P (a # list)) \ P list" by (rule list.induct) text \Compatibility:\ setup \Sign.mandatory_path "list"\ lemmas inducts = list.induct lemmas recs = list.rec lemmas cases = list.case setup \Sign.parent_path\ lemmas set_simps = list.set (* legacy *) syntax \ \list Enumeration\ "_list" :: "args => 'a list" ("[(_)]") translations "[x, xs]" == "x#[xs]" "[x]" == "x#[]" subsection \Basic list processing functions\ primrec (nonexhaustive) last :: "'a list \ 'a" where "last (x # xs) = (if xs = [] then x else last xs)" primrec butlast :: "'a list \ 'a list" where "butlast [] = []" | "butlast (x # xs) = (if xs = [] then [] else x # butlast xs)" lemma set_rec: "set xs = rec_list {} (\x _. insert x) xs" by (induct xs) auto definition coset :: "'a list \ 'a set" where [simp]: "coset xs = - set xs" primrec append :: "'a list \ 'a list \ 'a list" (infixr "@" 65) where append_Nil: "[] @ ys = ys" | append_Cons: "(x#xs) @ ys = x # xs @ ys" primrec rev :: "'a list \ 'a list" where "rev [] = []" | "rev (x # xs) = rev xs @ [x]" primrec filter:: "('a \ bool) \ 'a list \ 'a list" where "filter P [] = []" | "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" text \Special input syntax for filter:\ syntax (ASCII) "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])") syntax "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_\_ ./ _])") translations "[x<-xs . P]" \ "CONST filter (\x. P) xs" primrec fold :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where fold_Nil: "fold f [] = id" | fold_Cons: "fold f (x # xs) = fold f xs \ f x" primrec foldr :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where foldr_Nil: "foldr f [] = id" | foldr_Cons: "foldr f (x # xs) = f x \ foldr f xs" primrec foldl :: "('b \ 'a \ 'b) \ 'b \ 'a list \ 'b" where foldl_Nil: "foldl f a [] = a" | foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs" primrec concat:: "'a list list \ 'a list" where "concat [] = []" | "concat (x # xs) = x @ concat xs" primrec drop:: "nat \ 'a list \ 'a list" where drop_Nil: "drop n [] = []" | drop_Cons: "drop n (x # xs) = (case n of 0 \ x # xs | Suc m \ drop m xs)" \ \Warning: simpset does not contain this definition, but separate theorems for \n = 0\ and \n = Suc k\\ primrec take:: "nat \ 'a list \ 'a list" where take_Nil:"take n [] = []" | take_Cons: "take n (x # xs) = (case n of 0 \ [] | Suc m \ x # take m xs)" \ \Warning: simpset does not contain this definition, but separate theorems for \n = 0\ and \n = Suc k\\ primrec (nonexhaustive) nth :: "'a list => nat => 'a" (infixl "!" 100) where nth_Cons: "(x # xs) ! n = (case n of 0 \ x | Suc k \ xs ! k)" \ \Warning: simpset does not contain this definition, but separate theorems for \n = 0\ and \n = Suc k\\ primrec list_update :: "'a list \ nat \ 'a \ 'a list" where "list_update [] i v = []" | "list_update (x # xs) i v = (case i of 0 \ v # xs | Suc j \ x # list_update xs j v)" nonterminal lupdbinds and lupdbind syntax "_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)") "" :: "lupdbind => lupdbinds" ("_") "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _") "_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [1000,0] 900) translations "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs" "xs[i:=x]" == "CONST list_update xs i x" primrec takeWhile :: "('a \ bool) \ 'a list \ 'a list" where "takeWhile P [] = []" | "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])" primrec dropWhile :: "('a \ bool) \ 'a list \ 'a list" where "dropWhile P [] = []" | "dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)" primrec zip :: "'a list \ 'b list \ ('a \ 'b) list" where "zip xs [] = []" | zip_Cons: "zip xs (y # ys) = (case xs of [] \ [] | z # zs \ (z, y) # zip zs ys)" \ \Warning: simpset does not contain this definition, but separate theorems for \xs = []\ and \xs = z # zs\\ abbreviation map2 :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where "map2 f xs ys \ map (\(x,y). f x y) (zip xs ys)" primrec product :: "'a list \ 'b list \ ('a \ 'b) list" where "product [] _ = []" | "product (x#xs) ys = map (Pair x) ys @ product xs ys" hide_const (open) product primrec product_lists :: "'a list list \ 'a list list" where "product_lists [] = [[]]" | "product_lists (xs # xss) = concat (map (\x. map (Cons x) (product_lists xss)) xs)" primrec upt :: "nat \ nat \ nat list" ("(1[_.. j then [i.. 'a list \ 'a list" where "insert x xs = (if x \ set xs then xs else x # xs)" definition union :: "'a list \ 'a list \ 'a list" where "union = fold insert" hide_const (open) insert union hide_fact (open) insert_def union_def primrec find :: "('a \ bool) \ 'a list \ 'a option" where "find _ [] = None" | "find P (x#xs) = (if P x then Some x else find P xs)" text \In the context of multisets, \count_list\ is equivalent to \<^term>\count \ mset\ and it it advisable to use the latter.\ primrec count_list :: "'a list \ 'a \ nat" where "count_list [] y = 0" | "count_list (x#xs) y = (if x=y then count_list xs y + 1 else count_list xs y)" definition "extract" :: "('a \ bool) \ 'a list \ ('a list * 'a * 'a list) option" where "extract P xs = (case dropWhile (Not \ P) xs of [] \ None | y#ys \ Some(takeWhile (Not \ P) xs, y, ys))" hide_const (open) "extract" primrec those :: "'a option list \ 'a list option" where "those [] = Some []" | "those (x # xs) = (case x of None \ None | Some y \ map_option (Cons y) (those xs))" primrec remove1 :: "'a \ 'a list \ 'a list" where "remove1 x [] = []" | "remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)" primrec removeAll :: "'a \ 'a list \ 'a list" where "removeAll x [] = []" | "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)" primrec distinct :: "'a list \ bool" where "distinct [] \ True" | "distinct (x # xs) \ x \ set xs \ distinct xs" primrec remdups :: "'a list \ 'a list" where "remdups [] = []" | "remdups (x # xs) = (if x \ set xs then remdups xs else x # remdups xs)" fun remdups_adj :: "'a list \ 'a list" where "remdups_adj [] = []" | "remdups_adj [x] = [x]" | "remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))" primrec replicate :: "nat \ 'a \ 'a list" where replicate_0: "replicate 0 x = []" | replicate_Suc: "replicate (Suc n) x = x # replicate n x" text \ Function \size\ is overloaded for all datatypes. Users may refer to the list version as \length\.\ abbreviation length :: "'a list \ nat" where "length \ size" definition enumerate :: "nat \ 'a list \ (nat \ 'a) list" where enumerate_eq_zip: "enumerate n xs = zip [n.. 'a list" where "rotate1 [] = []" | "rotate1 (x # xs) = xs @ [x]" definition rotate :: "nat \ 'a list \ 'a list" where "rotate n = rotate1 ^^ n" definition nths :: "'a list => nat set => 'a list" where "nths xs A = map fst (filter (\p. snd p \ A) (zip xs [0.. 'a list list" where "subseqs [] = [[]]" | "subseqs (x#xs) = (let xss = subseqs xs in map (Cons x) xss @ xss)" primrec n_lists :: "nat \ 'a list \ 'a list list" where "n_lists 0 xs = [[]]" | "n_lists (Suc n) xs = concat (map (\ys. map (\y. y # ys) xs) (n_lists n xs))" hide_const (open) n_lists function splice :: "'a list \ 'a list \ 'a list" where "splice [] ys = ys" | "splice (x#xs) ys = x # splice ys xs" by pat_completeness auto termination by(relation "measure(\(xs,ys). size xs + size ys)") auto function shuffles where "shuffles [] ys = {ys}" | "shuffles xs [] = {xs}" | "shuffles (x # xs) (y # ys) = (#) x ` shuffles xs (y # ys) \ (#) y ` shuffles (x # xs) ys" by pat_completeness simp_all termination by lexicographic_order text\Use only if you cannot use \<^const>\Min\ instead:\ fun min_list :: "'a::ord list \ 'a" where "min_list (x # xs) = (case xs of [] \ x | _ \ min x (min_list xs))" text\Returns first minimum:\ fun arg_min_list :: "('a \ ('b::linorder)) \ 'a list \ 'a" where "arg_min_list f [x] = x" | "arg_min_list f (x#y#zs) = (let m = arg_min_list f (y#zs) in if f x \ f m then x else m)" text\ \begin{figure}[htbp] \fbox{ \begin{tabular}{l} @{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\ @{lemma "length [a,b,c] = 3" by simp}\\ @{lemma "set [a,b,c] = {a,b,c}" by simp}\\ @{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\ @{lemma "rev [a,b,c] = [c,b,a]" by simp}\\ @{lemma "hd [a,b,c,d] = a" by simp}\\ @{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\ @{lemma "last [a,b,c,d] = d" by simp}\\ @{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\ @{lemma[source] "filter (\n::nat. n<2) [0,2,1] = [0,1]" by simp}\\ @{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\ @{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\ @{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\ @{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\ @{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\ @{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\ @{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\ @{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\ @{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\ @{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\ @{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\ @{lemma "shuffles [a,b] [c,d] = {[a,b,c,d],[a,c,b,d],[a,c,d,b],[c,a,b,d],[c,a,d,b],[c,d,a,b]}" by (simp add: insert_commute)}\\ @{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\ @{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\ @{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\ @{lemma "drop 6 [a,b,c,d] = []" by simp}\\ @{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\ @{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\ @{lemma "distinct [2,0,1::nat]" by simp}\\ @{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\ @{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\ @{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\ @{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\ @{lemma "List.union [2,3,4] [0::int,1,2] = [4,3,0,1,2]" by (simp add: List.insert_def List.union_def)}\\ @{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\ @{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\ @{lemma "count_list [0,1,0,2::int] 0 = 2" by (simp)}\\ @{lemma "List.extract (%i::int. i>0) [0,0] = None" by(simp add: extract_def)}\\ @{lemma "List.extract (%i::int. i>0) [0,1,0,2] = Some([0], 1, [0,2])" by(simp add: extract_def)}\\ @{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\ @{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\ @{lemma "nth [a,b,c,d] 2 = c" by simp}\\ @{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\ @{lemma "nths [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:nths_def)}\\ @{lemma "subseqs [a,b] = [[a, b], [a], [b], []]" by simp}\\ @{lemma "List.n_lists 2 [a,b,c] = [[a, a], [b, a], [c, a], [a, b], [b, b], [c, b], [a, c], [b, c], [c, c]]" by (simp add: eval_nat_numeral)}\\ @{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\ @{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\ @{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\ @{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\ @{lemma "min_list [3,1,-2::int] = -2" by (simp)}\\ @{lemma "arg_min_list (\i. i*i) [3,-1,1,-2::int] = -1" by (simp)} \end{tabular}} \caption{Characteristic examples} \label{fig:Characteristic} \end{figure} Figure~\ref{fig:Characteristic} shows characteristic examples that should give an intuitive understanding of the above functions. \ text\The following simple sort functions are intended for proofs, not for efficient implementations.\ text \A sorted predicate w.r.t. a relation:\ fun sorted_wrt :: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted_wrt P [] = True" | "sorted_wrt P (x # ys) = ((\y \ set ys. P x y) \ sorted_wrt P ys)" (* FIXME: define sorted in terms of sorted_wrt *) text \A class-based sorted predicate:\ context linorder begin fun sorted :: "'a list \ bool" where "sorted [] = True" | "sorted (x # ys) = ((\y \ set ys. x \ y) \ sorted ys)" lemma sorted_sorted_wrt: "sorted = sorted_wrt (\)" proof (rule ext) fix xs show "sorted xs = sorted_wrt (\) xs" by(induction xs rule: sorted.induct) auto qed primrec insort_key :: "('b \ 'a) \ 'b \ 'b list \ 'b list" where "insort_key f x [] = [x]" | "insort_key f x (y#ys) = (if f x \ f y then (x#y#ys) else y#(insort_key f x ys))" definition sort_key :: "('b \ 'a) \ 'b list \ 'b list" where "sort_key f xs = foldr (insort_key f) xs []" definition insort_insert_key :: "('b \ 'a) \ 'b \ 'b list \ 'b list" where "insort_insert_key f x xs = (if f x \ f ` set xs then xs else insort_key f x xs)" abbreviation "sort \ sort_key (\x. x)" abbreviation "insort \ insort_key (\x. x)" abbreviation "insort_insert \ insort_insert_key (\x. x)" definition stable_sort_key :: "(('b \ 'a) \ 'b list \ 'b list) \ bool" where "stable_sort_key sk = (\f xs k. filter (\y. f y = k) (sk f xs) = filter (\y. f y = k) xs)" end subsubsection \List comprehension\ text\Input syntax for Haskell-like list comprehension notation. Typical example: \[(x,y). x \ xs, y \ ys, x \ y]\, the list of all pairs of distinct elements from \xs\ and \ys\. The syntax is as in Haskell, except that \|\ becomes a dot (like in Isabelle's set comprehension): \[e. x \ xs, \]\ rather than \verb![e| x <- xs, ...]!. The qualifiers after the dot are \begin{description} \item[generators] \p \ xs\, where \p\ is a pattern and \xs\ an expression of list type, or \item[guards] \b\, where \b\ is a boolean expression. %\item[local bindings] @ {text"let x = e"}. \end{description} Just like in Haskell, list comprehension is just a shorthand. To avoid misunderstandings, the translation into desugared form is not reversed upon output. Note that the translation of \[e. x \ xs]\ is optmized to \<^term>\map (%x. e) xs\. It is easy to write short list comprehensions which stand for complex expressions. During proofs, they may become unreadable (and mangled). In such cases it can be advisable to introduce separate definitions for the list comprehensions in question.\ nonterminal lc_qual and lc_quals syntax "_listcompr" :: "'a \ lc_qual \ lc_quals \ 'a list" ("[_ . __") "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ \ _") "_lc_test" :: "bool \ lc_qual" ("_") (*"_lc_let" :: "letbinds => lc_qual" ("let _")*) "_lc_end" :: "lc_quals" ("]") "_lc_quals" :: "lc_qual \ lc_quals \ lc_quals" (", __") syntax (ASCII) "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ <- _") parse_translation \ let val NilC = Syntax.const \<^const_syntax>\Nil\; val ConsC = Syntax.const \<^const_syntax>\Cons\; val mapC = Syntax.const \<^const_syntax>\map\; val concatC = Syntax.const \<^const_syntax>\concat\; val IfC = Syntax.const \<^const_syntax>\If\; val dummyC = Syntax.const \<^const_syntax>\Pure.dummy_pattern\ fun single x = ConsC $ x $ NilC; fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *) let (* FIXME proper name context!? *) val x = Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT); val e = if opti then single e else e; val case1 = Syntax.const \<^syntax_const>\_case1\ $ p $ e; val case2 = Syntax.const \<^syntax_const>\_case1\ $ dummyC $ NilC; val cs = Syntax.const \<^syntax_const>\_case2\ $ case1 $ case2; in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end; fun pair_pat_tr (x as Free _) e = Syntax_Trans.abs_tr [x, e] | pair_pat_tr (_ $ p1 $ p2) e = Syntax.const \<^const_syntax>\case_prod\ $ pair_pat_tr p1 (pair_pat_tr p2 e) | pair_pat_tr dummy e = Syntax_Trans.abs_tr [Syntax.const "_idtdummy", e] fun pair_pat ctxt (Const (\<^const_syntax>\Pair\,_) $ s $ t) = pair_pat ctxt s andalso pair_pat ctxt t | pair_pat ctxt (Free (s,_)) = let val thy = Proof_Context.theory_of ctxt; val s' = Proof_Context.intern_const ctxt s; in not (Sign.declared_const thy s') end | pair_pat _ t = (t = dummyC); fun abs_tr ctxt p e opti = let val p = Term_Position.strip_positions p in if pair_pat ctxt p then (pair_pat_tr p e, true) else (pat_tr ctxt p e opti, false) end fun lc_tr ctxt [e, Const (\<^syntax_const>\_lc_test\, _) $ b, qs] = let val res = (case qs of Const (\<^syntax_const>\_lc_end\, _) => single e | Const (\<^syntax_const>\_lc_quals\, _) $ q $ qs => lc_tr ctxt [e, q, qs]); in IfC $ b $ res $ NilC end | lc_tr ctxt [e, Const (\<^syntax_const>\_lc_gen\, _) $ p $ es, Const(\<^syntax_const>\_lc_end\, _)] = (case abs_tr ctxt p e true of (f, true) => mapC $ f $ es | (f, false) => concatC $ (mapC $ f $ es)) | lc_tr ctxt [e, Const (\<^syntax_const>\_lc_gen\, _) $ p $ es, Const (\<^syntax_const>\_lc_quals\, _) $ q $ qs] = let val e' = lc_tr ctxt [e, q, qs]; in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end; in [(\<^syntax_const>\_listcompr\, lc_tr)] end \ ML_val \ let val read = Syntax.read_term \<^context> o Syntax.implode_input; fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote (#1 (Input.source_content s1)) ^ Position.here_list [Input.pos_of s1, Input.pos_of s2]); in check \[(x,y,z). b]\ \if b then [(x, y, z)] else []\; check \[(x,y,z). (x,_,y)\xs]\ \map (\(x,_,y). (x, y, z)) xs\; check \[e x y. (x,_)\xs, y\ys]\ \concat (map (\(x,_). map (\y. e x y) ys) xs)\; check \[(x,y,z). xb]\ \if x < a then if b < x then [(x, y, z)] else [] else []\; check \[(x,y,z). x\xs, x>b]\ \concat (map (\x. if b < x then [(x, y, z)] else []) xs)\; check \[(x,y,z). xxs]\ \if x < a then map (\x. (x, y, z)) xs else []\; check \[(x,y). Cons True x \ xs]\ \concat (map (\xa. case xa of [] \ [] | True # x \ [(x, y)] | False # x \ []) xs)\; check \[(x,y,z). Cons x [] \ xs]\ \concat (map (\xa. case xa of [] \ [] | [x] \ [(x, y, z)] | x # aa # lista \ []) xs)\; check \[(x,y,z). xb, x=d]\ \if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []\; check \[(x,y,z). xb, y\ys]\ \if x < a then if b < x then map (\y. (x, y, z)) ys else [] else []\; check \[(x,y,z). xxs,y>b]\ \if x < a then concat (map (\(_,x). if b < y then [(x, y, z)] else []) xs) else []\; check \[(x,y,z). xxs, y\ys]\ \if x < a then concat (map (\x. map (\y. (x, y, z)) ys) xs) else []\; check \[(x,y,z). x\xs, x>b, y \concat (map (\x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)\; check \[(x,y,z). x\xs, x>b, y\ys]\ \concat (map (\x. if b < x then map (\y. (x, y, z)) ys else []) xs)\; check \[(x,y,z). x\xs, (y,_)\ys,y>x]\ \concat (map (\x. concat (map (\(y,_). if x < y then [(x, y, z)] else []) ys)) xs)\; check \[(x,y,z). x\xs, y\ys,z\zs]\ \concat (map (\x. concat (map (\y. map (\z. (x, y, z)) zs) ys)) xs)\ end; \ ML \ (* Simproc for rewriting list comprehensions applied to List.set to set comprehension. *) signature LIST_TO_SET_COMPREHENSION = sig val simproc : Proof.context -> cterm -> thm option end structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION = struct (* conversion *) fun all_exists_conv cv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Ex\, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct | _ => cv ctxt ct) fun all_but_last_exists_conv cv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Ex\, _) $ Abs (_, _, Const (\<^const_name>\Ex\, _) $ _) => Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct | _ => cv ctxt ct) fun Collect_conv cv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Collect\, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct | _ => raise CTERM ("Collect_conv", [ct])) fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th) fun conjunct_assoc_conv ct = Conv.try_conv (rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct fun right_hand_set_comprehension_conv conv ctxt = HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (all_exists_conv conv o #2) ctxt)) (* term abstraction of list comprehension patterns *) datatype termlets = If | Case of typ * int local val set_Nil_I = @{lemma "set [] = {x. False}" by (simp add: empty_def [symmetric])} val set_singleton = @{lemma "set [a] = {x. x = a}" by simp} val inst_Collect_mem_eq = @{lemma "set A = {x. x \ set A}" by simp} val del_refl_eq = @{lemma "(t = t \ P) \ P" by simp} fun mk_set T = Const (\<^const_name>\set\, HOLogic.listT T --> HOLogic.mk_setT T) fun dest_set (Const (\<^const_name>\set\, _) $ xs) = xs fun dest_singleton_list (Const (\<^const_name>\Cons\, _) $ t $ (Const (\<^const_name>\Nil\, _))) = t | dest_singleton_list t = raise TERM ("dest_singleton_list", [t]) (*We check that one case returns a singleton list and all other cases return [], and return the index of the one singleton list case.*) fun possible_index_of_singleton_case cases = let fun check (i, case_t) s = (case strip_abs_body case_t of (Const (\<^const_name>\Nil\, _)) => s | _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE)) in fold_index check cases (SOME NONE) |> the_default NONE end (*returns condition continuing term option*) fun dest_if (Const (\<^const_name>\If\, _) $ cond $ then_t $ Const (\<^const_name>\Nil\, _)) = SOME (cond, then_t) | dest_if _ = NONE (*returns (case_expr type index chosen_case constr_name) option*) fun dest_case ctxt case_term = let val (case_const, args) = strip_comb case_term in (case try dest_Const case_const of SOME (c, T) => (case Ctr_Sugar.ctr_sugar_of_case ctxt c of SOME {ctrs, ...} => (case possible_index_of_singleton_case (fst (split_last args)) of SOME i => let val constr_names = map (fst o dest_Const) ctrs val (Ts, _) = strip_type T val T' = List.last Ts in SOME (List.last args, T', i, nth args i, nth constr_names i) end | NONE => NONE) | NONE => NONE) | NONE => NONE) end fun tac ctxt [] = resolve_tac ctxt [set_singleton] 1 ORELSE resolve_tac ctxt [inst_Collect_mem_eq] 1 | tac ctxt (If :: cont) = Splitter.split_tac ctxt @{thms if_split} 1 THEN resolve_tac ctxt @{thms conjI} 1 THEN resolve_tac ctxt @{thms impI} 1 THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (right_hand_set_comprehension_conv (K (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv then_conv rewr_conv' @{lemma "(True \ P) = P" by simp})) ctxt') 1) ctxt 1 THEN tac ctxt cont THEN resolve_tac ctxt @{thms impI} 1 THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (right_hand_set_comprehension_conv (K (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv then_conv rewr_conv' @{lemma "(False \ P) = False" by simp})) ctxt') 1) ctxt 1 THEN resolve_tac ctxt [set_Nil_I] 1 | tac ctxt (Case (T, i) :: cont) = let val SOME {injects, distincts, case_thms, split, ...} = Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T)) in (* do case distinction *) Splitter.split_tac ctxt [split] 1 THEN EVERY (map_index (fn (i', _) => (if i' < length case_thms - 1 then resolve_tac ctxt @{thms conjI} 1 else all_tac) THEN REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1) THEN resolve_tac ctxt @{thms impI} 1 THEN (if i' = i then (* continue recursively *) Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K ((HOLogic.conj_conv (HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects)))) Conv.all_conv) then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq)) then_conv conjunct_assoc_conv)) ctxt' then_conv (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') => Conv.repeat_conv (all_but_last_exists_conv (K (rewr_conv' @{lemma "(\x. x = t \ P x) = P t" by simp})) ctxt'')) ctxt')))) 1) ctxt 1 THEN tac ctxt cont else Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (right_hand_set_comprehension_conv (K (HOLogic.conj_conv ((HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems))) then_conv (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts))) Conv.all_conv then_conv (rewr_conv' @{lemma "(False \ P) = False" by simp}))) ctxt' then_conv HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') => Conv.repeat_conv (Conv.bottom_conv (K (rewr_conv' @{lemma "(\x. P) = P" by simp})) ctxt'')) ctxt'))) 1) ctxt 1 THEN resolve_tac ctxt [set_Nil_I] 1)) case_thms) end in fun simproc ctxt redex = let fun make_inner_eqs bound_vs Tis eqs t = (case dest_case ctxt t of SOME (x, T, i, cont, constr_name) => let val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont) val x' = incr_boundvars (length vs) x val eqs' = map (incr_boundvars (length vs)) eqs val constr_t = list_comb (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0)) val constr_eq = Const (\<^const_name>\HOL.eq\, T --> T --> \<^typ>\bool\) $ constr_t $ x' in make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body end | NONE => (case dest_if t of SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont | NONE => if null eqs then NONE (*no rewriting, nothing to be done*) else let val Type (\<^type_name>\list\, [rT]) = fastype_of1 (map snd bound_vs, t) val pat_eq = (case try dest_singleton_list t of SOME t' => Const (\<^const_name>\HOL.eq\, rT --> rT --> \<^typ>\bool\) $ Bound (length bound_vs) $ t' | NONE => Const (\<^const_name>\Set.member\, rT --> HOLogic.mk_setT rT --> \<^typ>\bool\) $ Bound (length bound_vs) $ (mk_set rT $ t)) val reverse_bounds = curry subst_bounds ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)]) val eqs' = map reverse_bounds eqs val pat_eq' = reverse_bounds pat_eq val inner_t = fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t) (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq') val lhs = Thm.term_of redex val rhs = HOLogic.mk_Collect ("x", rT, inner_t) val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) in SOME ((Goal.prove ctxt [] [] rewrite_rule_t (fn {context = ctxt', ...} => tac ctxt' (rev Tis))) RS @{thm eq_reflection}) end)) in make_inner_eqs [] [] [] (dest_set (Thm.term_of redex)) end end end \ simproc_setup list_to_set_comprehension ("set xs") = \K List_to_Set_Comprehension.simproc\ code_datatype set coset hide_const (open) coset subsubsection \\<^const>\Nil\ and \<^const>\Cons\\ lemma not_Cons_self [simp]: "xs \ x # xs" by (induct xs) auto lemma not_Cons_self2 [simp]: "x # xs \ xs" by (rule not_Cons_self [symmetric]) lemma neq_Nil_conv: "(xs \ []) = (\y ys. xs = y # ys)" by (induct xs) auto lemma tl_Nil: "tl xs = [] \ xs = [] \ (\x. xs = [x])" by (cases xs) auto lemma Nil_tl: "[] = tl xs \ xs = [] \ (\x. xs = [x])" by (cases xs) auto lemma length_induct: "(\xs. \ys. length ys < length xs \ P ys \ P xs) \ P xs" by (fact measure_induct) lemma induct_list012: "\P []; \x. P [x]; \x y zs. \ P zs; P (y # zs) \ \ P (x # y # zs)\ \ P xs" by induction_schema (pat_completeness, lexicographic_order) lemma list_nonempty_induct [consumes 1, case_names single cons]: "\ xs \ []; \x. P [x]; \x xs. xs \ [] \ P xs \ P (x # xs)\ \ P xs" by(induction xs rule: induct_list012) auto lemma inj_split_Cons: "inj_on (\(xs, n). n#xs) X" by (auto intro!: inj_onI) lemma inj_on_Cons1 [simp]: "inj_on ((#) x) A" by(simp add: inj_on_def) subsubsection \\<^const>\length\\ text \ Needs to come before \@\ because of theorem \append_eq_append_conv\. \ lemma length_append [simp]: "length (xs @ ys) = length xs + length ys" by (induct xs) auto lemma length_map [simp]: "length (map f xs) = length xs" by (induct xs) auto lemma length_rev [simp]: "length (rev xs) = length xs" by (induct xs) auto lemma length_tl [simp]: "length (tl xs) = length xs - 1" by (cases xs) auto lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])" by (induct xs) auto lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \ [])" by (induct xs) auto lemma length_pos_if_in_set: "x \ set xs \ length xs > 0" by auto lemma length_Suc_conv: "(length xs = Suc n) = (\y ys. xs = y # ys \ length ys = n)" by (induct xs) auto lemma Suc_length_conv: "(Suc n = length xs) = (\y ys. xs = y # ys \ length ys = n)" by (induct xs; simp; blast) lemma Suc_le_length_iff: "(Suc n \ length xs) = (\x ys. xs = x # ys \ n \ length ys)" by (metis Suc_le_D[of n] Suc_le_mono[of n] Suc_length_conv[of _ xs]) lemma impossible_Cons: "length xs \ length ys ==> xs = x # ys = False" by (induct xs) auto lemma list_induct2 [consumes 1, case_names Nil Cons]: "length xs = length ys \ P [] [] \ (\x xs y ys. length xs = length ys \ P xs ys \ P (x#xs) (y#ys)) \ P xs ys" proof (induct xs arbitrary: ys) case (Cons x xs ys) then show ?case by (cases ys) simp_all qed simp lemma list_induct3 [consumes 2, case_names Nil Cons]: "length xs = length ys \ length ys = length zs \ P [] [] [] \ (\x xs y ys z zs. length xs = length ys \ length ys = length zs \ P xs ys zs \ P (x#xs) (y#ys) (z#zs)) \ P xs ys zs" proof (induct xs arbitrary: ys zs) case Nil then show ?case by simp next case (Cons x xs ys zs) then show ?case by (cases ys, simp_all) (cases zs, simp_all) qed lemma list_induct4 [consumes 3, case_names Nil Cons]: "length xs = length ys \ length ys = length zs \ length zs = length ws \ P [] [] [] [] \ (\x xs y ys z zs w ws. length xs = length ys \ length ys = length zs \ length zs = length ws \ P xs ys zs ws \ P (x#xs) (y#ys) (z#zs) (w#ws)) \ P xs ys zs ws" proof (induct xs arbitrary: ys zs ws) case Nil then show ?case by simp next case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all) qed lemma list_induct2': "\ P [] []; \x xs. P (x#xs) []; \y ys. P [] (y#ys); \x xs y ys. P xs ys \ P (x#xs) (y#ys) \ \ P xs ys" by (induct xs arbitrary: ys) (case_tac x, auto)+ lemma list_all2_iff: "list_all2 P xs ys \ length xs = length ys \ (\(x, y) \ set (zip xs ys). P x y)" by (induct xs ys rule: list_induct2') auto lemma neq_if_length_neq: "length xs \ length ys \ (xs = ys) == False" by (rule Eq_FalseI) auto simproc_setup list_neq ("(xs::'a list) = ys") = \ (* Reduces xs=ys to False if xs and ys cannot be of the same length. This is the case if the atomic sublists of one are a submultiset of those of the other list and there are fewer Cons's in one than the other. *) let fun len (Const(\<^const_name>\Nil\,_)) acc = acc | len (Const(\<^const_name>\Cons\,_) $ _ $ xs) (ts,n) = len xs (ts,n+1) | len (Const(\<^const_name>\append\,_) $ xs $ ys) acc = len xs (len ys acc) | len (Const(\<^const_name>\rev\,_) $ xs) acc = len xs acc | len (Const(\<^const_name>\map\,_) $ _ $ xs) acc = len xs acc | len t (ts,n) = (t::ts,n); val ss = simpset_of \<^context>; fun list_neq ctxt ct = let val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct; val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0); fun prove_neq() = let val Type(_,listT::_) = eqT; val size = HOLogic.size_const listT; val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs); val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len); val thm = Goal.prove ctxt [] [] neq_len (K (simp_tac (put_simpset ss ctxt) 1)); in SOME (thm RS @{thm neq_if_length_neq}) end in if m < n andalso submultiset (op aconv) (ls,rs) orelse n < m andalso submultiset (op aconv) (rs,ls) then prove_neq() else NONE end; in K list_neq end \ subsubsection \\@\ -- append\ global_interpretation append: monoid append Nil proof fix xs ys zs :: "'a list" show "(xs @ ys) @ zs = xs @ (ys @ zs)" by (induct xs) simp_all show "xs @ [] = xs" by (induct xs) simp_all qed simp lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)" by (fact append.assoc) lemma append_Nil2: "xs @ [] = xs" by (fact append.right_neutral) lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \ ys = [])" by (induct xs) auto lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \ ys = [])" by (induct xs) auto lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])" by (induct xs) auto lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])" by (induct xs) auto lemma append_eq_append_conv [simp]: "length xs = length ys \ length us = length vs ==> (xs@us = ys@vs) = (xs=ys \ us=vs)" by (induct xs arbitrary: ys; case_tac ys; force) lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) = (\us. xs = zs @ us \ us @ ys = ts \ xs @ us = zs \ ys = us @ ts)" proof (induct xs arbitrary: ys zs ts) case (Cons x xs) then show ?case by (cases zs) auto qed fastforce lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)" by simp lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \ x = y)" by simp lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)" by simp lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])" using append_same_eq [of _ _ "[]"] by auto lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])" using append_same_eq [of "[]"] by auto lemma hd_Cons_tl: "xs \ [] ==> hd xs # tl xs = xs" by (fact list.collapse) lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)" by (induct xs) auto lemma hd_append2 [simp]: "xs \ [] ==> hd (xs @ ys) = hd xs" by (simp add: hd_append split: list.split) lemma tl_append: "tl (xs @ ys) = (case xs of [] \ tl ys | z#zs \ zs @ ys)" by (simp split: list.split) lemma tl_append2 [simp]: "xs \ [] ==> tl (xs @ ys) = tl xs @ ys" by (simp add: tl_append split: list.split) lemma Cons_eq_append_conv: "x#xs = ys@zs = (ys = [] \ x#xs = zs \ (\ys'. x#ys' = ys \ xs = ys'@zs))" by(cases ys) auto lemma append_eq_Cons_conv: "(ys@zs = x#xs) = (ys = [] \ zs = x#xs \ (\ys'. ys = x#ys' \ ys'@zs = xs))" by(cases ys) auto lemma longest_common_prefix: "\ps xs' ys'. xs = ps @ xs' \ ys = ps @ ys' \ (xs' = [] \ ys' = [] \ hd xs' \ hd ys')" by (induct xs ys rule: list_induct2') (blast, blast, blast, metis (no_types, hide_lams) append_Cons append_Nil list.sel(1)) text \Trivial rules for solving \@\-equations automatically.\ lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys" by simp lemma Cons_eq_appendI: "[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs" by (drule sym) simp lemma append_eq_appendI: "[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us" by (drule sym) simp text \ Simplification procedure for all list equalities. Currently only tries to rearrange \@\ to see if - both lists end in a singleton list, - or both lists end in the same list. \ simproc_setup list_eq ("(xs::'a list) = ys") = \ let fun last (cons as Const (\<^const_name>\Cons\, _) $ _ $ xs) = (case xs of Const (\<^const_name>\Nil\, _) => cons | _ => last xs) | last (Const(\<^const_name>\append\,_) $ _ $ ys) = last ys | last t = t; fun list1 (Const(\<^const_name>\Cons\,_) $ _ $ Const(\<^const_name>\Nil\,_)) = true | list1 _ = false; fun butlast ((cons as Const(\<^const_name>\Cons\,_) $ x) $ xs) = (case xs of Const (\<^const_name>\Nil\, _) => xs | _ => cons $ butlast xs) | butlast ((app as Const (\<^const_name>\append\, _) $ xs) $ ys) = app $ butlast ys | butlast xs = Const(\<^const_name>\Nil\, fastype_of xs); val rearr_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]); fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) = let val lastl = last lhs and lastr = last rhs; fun rearr conv = let val lhs1 = butlast lhs and rhs1 = butlast rhs; val Type(_,listT::_) = eqT val appT = [listT,listT] ---> listT val app = Const(\<^const_name>\append\,appT) val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2)); val thm = Goal.prove ctxt [] [] eq (K (simp_tac (put_simpset rearr_ss ctxt) 1)); in SOME ((conv RS (thm RS trans)) RS eq_reflection) end; in if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv} else if lastl aconv lastr then rearr @{thm append_same_eq} else NONE end; in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end \ subsubsection \\<^const>\map\\ lemma hd_map: "xs \ [] \ hd (map f xs) = f (hd xs)" by (cases xs) simp_all lemma map_tl: "map f (tl xs) = tl (map f xs)" by (cases xs) simp_all lemma map_ext: "(\x. x \ set xs \ f x = g x) ==> map f xs = map g xs" by (induct xs) simp_all lemma map_ident [simp]: "map (\x. x) = (\xs. xs)" by (rule ext, induct_tac xs) auto lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys" by (induct xs) auto lemma map_map [simp]: "map f (map g xs) = map (f \ g) xs" by (induct xs) auto lemma map_comp_map[simp]: "((map f) \ (map g)) = map(f \ g)" by (rule ext) simp lemma rev_map: "rev (map f xs) = map f (rev xs)" by (induct xs) auto lemma map_eq_conv[simp]: "(map f xs = map g xs) = (\x \ set xs. f x = g x)" by (induct xs) auto lemma map_cong [fundef_cong]: "xs = ys \ (\x. x \ set ys \ f x = g x) \ map f xs = map g ys" by simp lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])" by (cases xs) auto lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])" by (cases xs) auto lemma map_eq_Cons_conv: "(map f xs = y#ys) = (\z zs. xs = z#zs \ f z = y \ map f zs = ys)" by (cases xs) auto lemma Cons_eq_map_conv: "(x#xs = map f ys) = (\z zs. ys = z#zs \ x = f z \ xs = map f zs)" by (cases ys) auto lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1] lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1] declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] lemma ex_map_conv: "(\xs. ys = map f xs) = (\y \ set ys. \x. y = f x)" by(induct ys, auto simp add: Cons_eq_map_conv) lemma map_eq_imp_length_eq: assumes "map f xs = map g ys" shows "length xs = length ys" using assms proof (induct ys arbitrary: xs) case Nil then show ?case by simp next case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto from Cons xs have "map f zs = map g ys" by simp with Cons have "length zs = length ys" by blast with xs show ?case by simp qed lemma map_inj_on: assumes map: "map f xs = map f ys" and inj: "inj_on f (set xs Un set ys)" shows "xs = ys" using map_eq_imp_length_eq [OF map] assms proof (induct rule: list_induct2) case (Cons x xs y ys) then show ?case by (auto intro: sym) qed auto lemma inj_on_map_eq_map: "inj_on f (set xs Un set ys) \ (map f xs = map f ys) = (xs = ys)" by(blast dest:map_inj_on) lemma map_injective: "map f xs = map f ys ==> inj f ==> xs = ys" by (induct ys arbitrary: xs) (auto dest!:injD) lemma inj_map_eq_map[simp]: "inj f \ (map f xs = map f ys) = (xs = ys)" by(blast dest:map_injective) lemma inj_mapI: "inj f ==> inj (map f)" by (iprover dest: map_injective injD intro: inj_onI) lemma inj_mapD: "inj (map f) ==> inj f" by (metis (no_types, hide_lams) injI list.inject list.simps(9) the_inv_f_f) lemma inj_map[iff]: "inj (map f) = inj f" by (blast dest: inj_mapD intro: inj_mapI) lemma inj_on_mapI: "inj_on f (\(set ` A)) \ inj_on (map f) A" by (blast intro:inj_onI dest:inj_onD map_inj_on) lemma map_idI: "(\x. x \ set xs \ f x = x) \ map f xs = xs" by (induct xs, auto) lemma map_fun_upd [simp]: "y \ set xs \ map (f(y:=v)) xs = map f xs" by (induct xs) auto lemma map_fst_zip[simp]: "length xs = length ys \ map fst (zip xs ys) = xs" by (induct rule:list_induct2, simp_all) lemma map_snd_zip[simp]: "length xs = length ys \ map snd (zip xs ys) = ys" by (induct rule:list_induct2, simp_all) lemma map_fst_zip_take: "map fst (zip xs ys) = take (min (length xs) (length ys)) xs" by (induct xs ys rule: list_induct2') simp_all lemma map_snd_zip_take: "map snd (zip xs ys) = take (min (length xs) (length ys)) ys" by (induct xs ys rule: list_induct2') simp_all lemma map2_map_map: "map2 h (map f xs) (map g xs) = map (\x. h (f x) (g x)) xs" by (induction xs) (auto) functor map: map by (simp_all add: id_def) declare map.id [simp] subsubsection \\<^const>\rev\\ lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs" by (induct xs) auto lemma rev_rev_ident [simp]: "rev (rev xs) = xs" by (induct xs) auto lemma rev_swap: "(rev xs = ys) = (xs = rev ys)" by auto lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])" by (induct xs) auto lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])" by (induct xs) auto lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])" by (cases xs) auto lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])" by (cases xs) auto lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by force next case Cons then show ?case by (cases ys) auto qed lemma inj_on_rev[iff]: "inj_on rev A" by(simp add:inj_on_def) lemma rev_induct [case_names Nil snoc]: "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs" apply(simplesubst rev_rev_ident[symmetric]) apply(rule_tac list = "rev xs" in list.induct, simp_all) done lemma rev_exhaust [case_names Nil snoc]: "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P" by (induct xs rule: rev_induct) auto lemmas rev_cases = rev_exhaust lemma rev_nonempty_induct [consumes 1, case_names single snoc]: assumes "xs \ []" and single: "\x. P [x]" and snoc': "\x xs. xs \ [] \ P xs \ P (xs@[x])" shows "P xs" using \xs \ []\ proof (induct xs rule: rev_induct) case (snoc x xs) then show ?case proof (cases xs) case Nil thus ?thesis by (simp add: single) next case Cons with snoc show ?thesis by (fastforce intro!: snoc') qed qed simp lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])" by(rule rev_cases[of xs]) auto subsubsection \\<^const>\set\\ declare list.set[code_post] \ \pretty output\ lemma finite_set [iff]: "finite (set xs)" by (induct xs) auto lemma set_append [simp]: "set (xs @ ys) = (set xs \ set ys)" by (induct xs) auto lemma hd_in_set[simp]: "xs \ [] \ hd xs \ set xs" by(cases xs) auto lemma set_subset_Cons: "set xs \ set (x # xs)" by auto lemma set_ConsD: "y \ set (x # xs) \ y=x \ y \ set xs" by auto lemma set_empty [iff]: "(set xs = {}) = (xs = [])" by (induct xs) auto lemma set_empty2[iff]: "({} = set xs) = (xs = [])" by(induct xs) auto lemma set_rev [simp]: "set (rev xs) = set xs" by (induct xs) auto lemma set_map [simp]: "set (map f xs) = f`(set xs)" by (induct xs) auto lemma set_filter [simp]: "set (filter P xs) = {x. x \ set xs \ P x}" by (induct xs) auto lemma set_upt [simp]: "set[i.. set xs \ \ys zs. xs = ys @ x # zs" proof (induct xs) case Nil thus ?case by simp next case Cons thus ?case by (auto intro: Cons_eq_appendI) qed lemma in_set_conv_decomp: "x \ set xs \ (\ys zs. xs = ys @ x # zs)" by (auto elim: split_list) lemma split_list_first: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set ys" proof (induct xs) case Nil thus ?case by simp next case (Cons a xs) show ?case proof cases assume "x = a" thus ?case using Cons by fastforce next assume "x \ a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI) qed qed lemma in_set_conv_decomp_first: "(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set ys)" by (auto dest!: split_list_first) lemma split_list_last: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set zs" proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc a xs) show ?case proof cases assume "x = a" thus ?case using snoc by (auto intro!: exI) next assume "x \ a" thus ?case using snoc by fastforce qed qed lemma in_set_conv_decomp_last: "(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set zs)" by (auto dest!: split_list_last) lemma split_list_prop: "\x \ set xs. P x \ \ys x zs. xs = ys @ x # zs \ P x" proof (induct xs) case Nil thus ?case by simp next case Cons thus ?case by(simp add:Bex_def)(metis append_Cons append.simps(1)) qed lemma split_list_propE: assumes "\x \ set xs. P x" obtains ys x zs where "xs = ys @ x # zs" and "P x" using split_list_prop [OF assms] by blast lemma split_list_first_prop: "\x \ set xs. P x \ \ys x zs. xs = ys@x#zs \ P x \ (\y \ set ys. \ P y)" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) show ?case proof cases assume "P x" hence "x # xs = [] @ x # xs \ P x \ (\y\set []. \ P y)" by simp thus ?thesis by fast next assume "\ P x" hence "\x\set xs. P x" using Cons(2) by simp thus ?thesis using \\ P x\ Cons(1) by (metis append_Cons set_ConsD) qed qed lemma split_list_first_propE: assumes "\x \ set xs. P x" obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\y \ set ys. \ P y" using split_list_first_prop [OF assms] by blast lemma split_list_first_prop_iff: "(\x \ set xs. P x) \ (\ys x zs. xs = ys@x#zs \ P x \ (\y \ set ys. \ P y))" by (rule, erule split_list_first_prop) auto lemma split_list_last_prop: "\x \ set xs. P x \ \ys x zs. xs = ys@x#zs \ P x \ (\z \ set zs. \ P z)" proof(induct xs rule:rev_induct) case Nil thus ?case by simp next case (snoc x xs) show ?case proof cases assume "P x" thus ?thesis by (auto intro!: exI) next assume "\ P x" hence "\x\set xs. P x" using snoc(2) by simp thus ?thesis using \\ P x\ snoc(1) by fastforce qed qed lemma split_list_last_propE: assumes "\x \ set xs. P x" obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\z \ set zs. \ P z" using split_list_last_prop [OF assms] by blast lemma split_list_last_prop_iff: "(\x \ set xs. P x) \ (\ys x zs. xs = ys@x#zs \ P x \ (\z \ set zs. \ P z))" by rule (erule split_list_last_prop, auto) lemma finite_list: "finite A \ \xs. set xs = A" by (erule finite_induct) (auto simp add: list.set(2)[symmetric] simp del: list.set(2)) lemma card_length: "card (set xs) \ length xs" by (induct xs) (auto simp add: card_insert_if) lemma set_minus_filter_out: "set xs - {y} = set (filter (\x. \ (x = y)) xs)" by (induct xs) auto lemma append_Cons_eq_iff: "\ x \ set xs; x \ set ys \ \ xs @ x # ys = xs' @ x # ys' \ (xs = xs' \ ys = ys')" by(auto simp: append_eq_Cons_conv Cons_eq_append_conv append_eq_append_conv2) subsubsection \\<^const>\filter\\ lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys" by (induct xs) auto lemma rev_filter: "rev (filter P xs) = filter P (rev xs)" by (induct xs) simp_all lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\x. Q x \ P x) xs" by (induct xs) auto lemma length_filter_le [simp]: "length (filter P xs) \ length xs" by (induct xs) (auto simp add: le_SucI) lemma sum_length_filter_compl: "length(filter P xs) + length(filter (\x. \P x) xs) = length xs" by(induct xs) simp_all lemma filter_True [simp]: "\x \ set xs. P x ==> filter P xs = xs" by (induct xs) auto lemma filter_False [simp]: "\x \ set xs. \ P x ==> filter P xs = []" by (induct xs) auto lemma filter_empty_conv: "(filter P xs = []) = (\x\set xs. \ P x)" by (induct xs) simp_all lemma filter_id_conv: "(filter P xs = xs) = (\x\set xs. P x)" proof (induct xs) case (Cons x xs) then show ?case using length_filter_le by (simp add: impossible_Cons) qed auto lemma filter_map: "filter P (map f xs) = map f (filter (P \ f) xs)" by (induct xs) simp_all lemma length_filter_map[simp]: "length (filter P (map f xs)) = length(filter (P \ f) xs)" by (simp add:filter_map) lemma filter_is_subset [simp]: "set (filter P xs) \ set xs" by auto lemma length_filter_less: "\ x \ set xs; \ P x \ \ length(filter P xs) < length xs" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) thus ?case using Suc_le_eq by fastforce qed lemma length_filter_conv_card: "length(filter p xs) = card{i. i < length xs \ p(xs!i)}" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) let ?S = "{i. i < length xs \ p(xs!i)}" have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite) show ?case (is "?l = card ?S'") proof (cases) assume "p x" hence eq: "?S' = insert 0 (Suc ` ?S)" by(auto simp: image_def split:nat.split dest:gr0_implies_Suc) have "length (filter p (x # xs)) = Suc(card ?S)" using Cons \p x\ by simp also have "\ = Suc(card(Suc ` ?S))" using fin by (simp add: card_image) also have "\ = card ?S'" using eq fin - by (simp add:card_insert_if) + by (simp add:card_insert_if) finally show ?thesis . next assume "\ p x" hence eq: "?S' = Suc ` ?S" by(auto simp add: image_def split:nat.split elim:lessE) have "length (filter p (x # xs)) = card ?S" using Cons \\ p x\ by simp also have "\ = card(Suc ` ?S)" using fin by (simp add: card_image) also have "\ = card ?S'" using eq fin by (simp add:card_insert_if) finally show ?thesis . qed qed lemma Cons_eq_filterD: "x#xs = filter P ys \ \us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs" (is "_ \ \us vs. ?P ys us vs") proof(induct ys) case Nil thus ?case by simp next case (Cons y ys) show ?case (is "\x. ?Q x") proof cases assume Py: "P y" show ?thesis proof cases assume "x = y" with Py Cons.prems have "?Q []" by simp then show ?thesis .. next assume "x \ y" with Py Cons.prems show ?thesis by simp qed next assume "\ P y" with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce then have "?Q (y#us)" by simp then show ?thesis .. qed qed lemma filter_eq_ConsD: "filter P ys = x#xs \ \us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs" by(rule Cons_eq_filterD) simp lemma filter_eq_Cons_iff: "(filter P ys = x#xs) = (\us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs)" by(auto dest:filter_eq_ConsD) lemma Cons_eq_filter_iff: "(x#xs = filter P ys) = (\us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs)" by(auto dest:Cons_eq_filterD) lemma inj_on_filter_key_eq: assumes "inj_on f (insert y (set xs))" shows "filter (\x. f y = f x) xs = filter (HOL.eq y) xs" using assms by (induct xs) auto lemma filter_cong[fundef_cong]: "xs = ys \ (\x. x \ set ys \ P x = Q x) \ filter P xs = filter Q ys" by (induct ys arbitrary: xs) auto subsubsection \List partitioning\ primrec partition :: "('a \ bool) \'a list \ 'a list \ 'a list" where "partition P [] = ([], [])" | "partition P (x # xs) = (let (yes, no) = partition P xs in if P x then (x # yes, no) else (yes, x # no))" lemma partition_filter1: "fst (partition P xs) = filter P xs" by (induct xs) (auto simp add: Let_def split_def) lemma partition_filter2: "snd (partition P xs) = filter (Not \ P) xs" by (induct xs) (auto simp add: Let_def split_def) lemma partition_P: assumes "partition P xs = (yes, no)" shows "(\p \ set yes. P p) \ (\p \ set no. \ P p)" proof - from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)" by simp_all then show ?thesis by (simp_all add: partition_filter1 partition_filter2) qed lemma partition_set: assumes "partition P xs = (yes, no)" shows "set yes \ set no = set xs" proof - from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)" by simp_all then show ?thesis by (auto simp add: partition_filter1 partition_filter2) qed lemma partition_filter_conv[simp]: "partition f xs = (filter f xs,filter (Not \ f) xs)" unfolding partition_filter2[symmetric] unfolding partition_filter1[symmetric] by simp declare partition.simps[simp del] subsubsection \\<^const>\concat\\ lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys" by (induct xs) auto lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\xs \ set xss. xs = [])" by (induct xss) auto lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\xs \ set xss. xs = [])" by (induct xss) auto lemma set_concat [simp]: "set (concat xs) = (\x\set xs. set x)" by (induct xs) auto lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs" by (induct xs) auto lemma map_concat: "map f (concat xs) = concat (map (map f) xs)" by (induct xs) auto lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)" by (induct xs) auto lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))" by (induct xs) auto lemma concat_eq_concat_iff: "\(x, y) \ set (zip xs ys). length x = length y ==> length xs = length ys ==> (concat xs = concat ys) = (xs = ys)" proof (induct xs arbitrary: ys) case (Cons x xs ys) thus ?case by (cases ys) auto qed (auto) lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \(x, y) \ set (zip xs ys). length x = length y ==> xs = ys" by (simp add: concat_eq_concat_iff) +lemma concat_eq_appendD: + assumes "concat xss = ys @ zs" "xss \ []" + shows "\xss1 xs xs' xss2. xss = xss1 @ (xs @ xs') # xss2 \ ys = concat xss1 @ xs \ zs = xs' @ concat xss2" + using assms +proof(induction xss arbitrary: ys) + case (Cons xs xss) + from Cons.prems consider + us where "xs @ us = ys" "concat xss = us @ zs" | + us where "xs = ys @ us" "us @ concat xss = zs" + by(auto simp add: append_eq_append_conv2) + then show ?case + proof cases + case 1 + then show ?thesis using Cons.IH[OF 1(2)] + by(cases xss)(auto intro: exI[where x="[]"], metis append.assoc append_Cons concat.simps(2)) + qed(auto intro: exI[where x="[]"]) +qed simp + +lemma concat_eq_append_conv: + "concat xss = ys @ zs \ + (if xss = [] then ys = [] \ zs = [] + else \xss1 xs xs' xss2. xss = xss1 @ (xs @ xs') # xss2 \ ys = concat xss1 @ xs \ zs = xs' @ concat xss2)" + by(auto dest: concat_eq_appendD) + subsubsection \\<^const>\nth\\ lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x" by auto lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n" by auto declare nth.simps [simp del] lemma nth_Cons_pos[simp]: "0 < n \ (x#xs) ! n = xs ! (n - 1)" by(auto simp: Nat.gr0_conv_Suc) lemma nth_append: "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))" proof (induct xs arbitrary: n) case (Cons x xs) then show ?case using less_Suc_eq_0_disj by auto qed simp lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x" by (induct xs) auto lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n" by (induct xs) auto lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)" proof (induct xs arbitrary: n) case (Cons x xs) then show ?case using less_Suc_eq_0_disj by auto qed simp lemma nth_tl: "n < length (tl xs) \ tl xs ! n = xs ! Suc n" by (induction xs) auto lemma hd_conv_nth: "xs \ [] \ hd xs = xs!0" by(cases xs) simp_all lemma list_eq_iff_nth_eq: "(xs = ys) = (length xs = length ys \ (\i ?R" by force show "?R \ ?L" using less_Suc_eq_0_disj by auto qed with Cons show ?case by simp qed simp lemma in_set_conv_nth: "(x \ set xs) = (\i < length xs. xs!i = x)" by(auto simp:set_conv_nth) lemma nth_equal_first_eq: assumes "x \ set xs" assumes "n \ length xs" shows "(x # xs) ! n = x \ n = 0" (is "?lhs \ ?rhs") proof assume ?lhs show ?rhs proof (rule ccontr) assume "n \ 0" then have "n > 0" by simp with \?lhs\ have "xs ! (n - 1) = x" by simp moreover from \n > 0\ \n \ length xs\ have "n - 1 < length xs" by simp ultimately have "\ix \ set xs\ in_set_conv_nth [of x xs] show False by simp qed next assume ?rhs then show ?lhs by simp qed lemma nth_non_equal_first_eq: assumes "x \ y" shows "(x # xs) ! n = y \ xs ! (n - 1) = y \ n > 0" (is "?lhs \ ?rhs") proof assume "?lhs" with assms have "n > 0" by (cases n) simp_all with \?lhs\ show ?rhs by simp next assume "?rhs" then show "?lhs" by simp qed lemma list_ball_nth: "\n < length xs; \x \ set xs. P x\ \ P(xs!n)" by (auto simp add: set_conv_nth) lemma nth_mem [simp]: "n < length xs \ xs!n \ set xs" by (auto simp add: set_conv_nth) lemma all_nth_imp_all_set: "\\i < length xs. P(xs!i); x \ set xs\ \ P x" by (auto simp add: set_conv_nth) lemma all_set_conv_all_nth: "(\x \ set xs. P x) = (\i. i < length xs \ P (xs ! i))" by (auto simp add: set_conv_nth) lemma rev_nth: "n < size xs \ rev xs ! n = xs ! (length xs - Suc n)" proof (induct xs arbitrary: n) case Nil thus ?case by simp next case (Cons x xs) hence n: "n < Suc (length xs)" by simp moreover { assume "n < length xs" with n obtain n' where n': "length xs - n = Suc n'" by (cases "length xs - n", auto) moreover from n' have "length xs - Suc n = n'" by simp ultimately have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp } ultimately show ?case by (clarsimp simp add: Cons nth_append) qed lemma Skolem_list_nth: "(\ix. P i x) = (\xs. size xs = k \ (\ixs. ?P k xs)") proof(induct k) case 0 show ?case by simp next case (Suc k) show ?case (is "?L = ?R" is "_ = (\xs. ?P' xs)") proof assume "?R" thus "?L" using Suc by auto next assume "?L" with Suc obtain x xs where "?P k xs \ P k x" by (metis less_Suc_eq) hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq) thus "?R" .. qed qed subsubsection \\<^const>\list_update\\ lemma length_list_update [simp]: "length(xs[i:=x]) = length xs" by (induct xs arbitrary: i) (auto split: nat.split) lemma nth_list_update: "i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)" by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x" by (simp add: nth_list_update) lemma nth_list_update_neq [simp]: "i \ j ==> xs[i:=x]!j = xs!j" by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) lemma list_update_id[simp]: "xs[i := xs!i] = xs" by (induct xs arbitrary: i) (simp_all split:nat.splits) lemma list_update_beyond[simp]: "length xs \ i \ xs[i:=x] = xs" proof (induct xs arbitrary: i) case (Cons x xs i) then show ?case by (metis leD length_list_update list_eq_iff_nth_eq nth_list_update_neq) qed simp lemma list_update_nonempty[simp]: "xs[k:=x] = [] \ xs=[]" by (simp only: length_0_conv[symmetric] length_list_update) lemma list_update_same_conv: "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)" by (induct xs arbitrary: i) (auto split: nat.split) lemma list_update_append1: "i < size xs \ (xs @ ys)[i:=x] = xs[i:=x] @ ys" by (induct xs arbitrary: i)(auto split:nat.split) lemma list_update_append: "(xs @ ys) [n:= x] = (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))" by (induct xs arbitrary: n) (auto split:nat.splits) lemma list_update_length [simp]: "(xs @ x # ys)[length xs := y] = (xs @ y # ys)" by (induct xs, auto) lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]" by(induct xs arbitrary: k)(auto split:nat.splits) lemma rev_update: "k < length xs \ rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]" by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits) lemma update_zip: "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])" by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split) lemma set_update_subset_insert: "set(xs[i:=x]) \ insert x (set xs)" by (induct xs arbitrary: i) (auto split: nat.split) lemma set_update_subsetI: "\set xs \ A; x \ A\ \ set(xs[i := x]) \ A" by (blast dest!: set_update_subset_insert [THEN subsetD]) lemma set_update_memI: "n < length xs \ x \ set (xs[n := x])" by (induct xs arbitrary: n) (auto split:nat.splits) lemma list_update_overwrite[simp]: "xs [i := x, i := y] = xs [i := y]" by (induct xs arbitrary: i) (simp_all split: nat.split) lemma list_update_swap: "i \ i' \ xs [i := x, i' := x'] = xs [i' := x', i := x]" by (induct xs arbitrary: i i') (simp_all split: nat.split) lemma list_update_code [code]: "[][i := y] = []" "(x # xs)[0 := y] = y # xs" "(x # xs)[Suc i := y] = x # xs[i := y]" by simp_all subsubsection \\<^const>\last\ and \<^const>\butlast\\ lemma last_snoc [simp]: "last (xs @ [x]) = x" by (induct xs) auto lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs" by (induct xs) auto lemma last_ConsL: "xs = [] \ last(x#xs) = x" by simp lemma last_ConsR: "xs \ [] \ last(x#xs) = last xs" by simp lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)" by (induct xs) (auto) lemma last_appendL[simp]: "ys = [] \ last(xs @ ys) = last xs" by(simp add:last_append) lemma last_appendR[simp]: "ys \ [] \ last(xs @ ys) = last ys" by(simp add:last_append) lemma last_tl: "xs = [] \ tl xs \ [] \last (tl xs) = last xs" by (induct xs) simp_all lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)" by (induct xs) simp_all lemma hd_rev: "xs \ [] \ hd(rev xs) = last xs" by(rule rev_exhaust[of xs]) simp_all lemma last_rev: "xs \ [] \ last(rev xs) = hd xs" by(cases xs) simp_all lemma last_in_set[simp]: "as \ [] \ last as \ set as" by (induct as) auto lemma length_butlast [simp]: "length (butlast xs) = length xs - 1" by (induct xs rule: rev_induct) auto lemma butlast_append: "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)" by (induct xs arbitrary: ys) auto lemma append_butlast_last_id [simp]: "xs \ [] \ butlast xs @ [last xs] = xs" by (induct xs) auto lemma in_set_butlastD: "x \ set (butlast xs) \ x \ set xs" by (induct xs) (auto split: if_split_asm) lemma in_set_butlast_appendI: "x \ set (butlast xs) \ x \ set (butlast ys) \ x \ set (butlast (xs @ ys))" by (auto dest: in_set_butlastD simp add: butlast_append) lemma last_drop[simp]: "n < length xs \ last (drop n xs) = last xs" by (induct xs arbitrary: n)(auto split:nat.split) lemma nth_butlast: assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n" proof (cases xs) case (Cons y ys) moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n" by (simp add: nth_append) ultimately show ?thesis using append_butlast_last_id by simp qed simp lemma last_conv_nth: "xs\[] \ last xs = xs!(length xs - 1)" by(induct xs)(auto simp:neq_Nil_conv) lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs" by (induction xs rule: induct_list012) simp_all lemma last_list_update: "xs \ [] \ last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)" by (auto simp: last_conv_nth) lemma butlast_list_update: "butlast(xs[k:=x]) = (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])" by(cases xs rule:rev_cases)(auto simp: list_update_append split: nat.splits) lemma last_map: "xs \ [] \ last (map f xs) = f (last xs)" by (cases xs rule: rev_cases) simp_all lemma map_butlast: "map f (butlast xs) = butlast (map f xs)" by (induct xs) simp_all lemma snoc_eq_iff_butlast: "xs @ [x] = ys \ (ys \ [] \ butlast ys = xs \ last ys = x)" by fastforce corollary longest_common_suffix: "\ss xs' ys'. xs = xs' @ ss \ ys = ys' @ ss \ (xs' = [] \ ys' = [] \ last xs' \ last ys')" using longest_common_prefix[of "rev xs" "rev ys"] unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv) lemma butlast_rev [simp]: "butlast (rev xs) = rev (tl xs)" by (cases xs) simp_all subsubsection \\<^const>\take\ and \<^const>\drop\\ lemma take_0: "take 0 xs = []" by (induct xs) auto lemma drop_0: "drop 0 xs = xs" by (induct xs) auto lemma take0[simp]: "take 0 = (\xs. [])" by(rule ext) (rule take_0) lemma drop0[simp]: "drop 0 = (\x. x)" by(rule ext) (rule drop_0) lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs" by simp lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs" by simp declare take_Cons [simp del] and drop_Cons [simp del] lemma take_Suc: "xs \ [] \ take (Suc n) xs = hd xs # take n (tl xs)" by(clarsimp simp add:neq_Nil_conv) lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)" by(cases xs, simp_all) lemma hd_take[simp]: "j > 0 \ hd (take j xs) = hd xs" by (metis gr0_conv_Suc list.sel(1) take.simps(1) take_Suc) lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)" by (induct xs arbitrary: n) simp_all lemma drop_tl: "drop n (tl xs) = tl(drop n xs)" by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split) lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)" by (cases n, simp, cases xs, auto) lemma tl_drop: "tl (drop n xs) = drop n (tl xs)" by (simp only: drop_tl) lemma nth_via_drop: "drop n xs = y#ys \ xs!n = y" by (induct xs arbitrary: n, simp)(auto simp: drop_Cons nth_Cons split: nat.splits) lemma take_Suc_conv_app_nth: "i < length xs \ take (Suc i) xs = take i xs @ [xs!i]" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma Cons_nth_drop_Suc: "i < length xs \ (xs!i) # (drop (Suc i) xs) = drop i xs" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma length_take [simp]: "length (take n xs) = min (length xs) n" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma length_drop [simp]: "length (drop n xs) = (length xs - n)" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_all [simp]: "length xs \ n ==> take n xs = xs" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma drop_all [simp]: "length xs \ n ==> drop n xs = []" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_append [simp]: "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma drop_append [simp]: "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_take [simp]: "take n (take m xs) = take (min n m) xs" proof (induct m arbitrary: xs n) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases n) simp_all qed lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs" proof (induct m arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)" proof (induct m arbitrary: xs n) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases n) simp_all qed lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)" by(induct xs arbitrary: m n)(auto simp: take_Cons drop_Cons split: nat.split) lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \ xs = [])" by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split) lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs \ n)" by (induct xs arbitrary: n) (auto simp: drop_Cons split:nat.split) lemma take_map: "take n (map f xs) = map f (take n xs)" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma drop_map: "drop n (map f xs) = map f (drop n xs)" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma drop_rev: "drop n (rev xs) = rev (take (length xs - n) xs)" by (cases "length xs < n") (auto simp: rev_take) lemma take_rev: "take n (rev xs) = rev (drop (length xs - n) xs)" by (cases "length xs < n") (auto simp: rev_drop) lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i" proof (induct xs arbitrary: i n) case Nil then show ?case by simp next case Cons then show ?case by (cases n; cases i) simp_all qed lemma nth_drop [simp]: "n \ length xs ==> (drop n xs)!i = xs!(n + i)" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma butlast_take: "n \ length xs ==> butlast (take n xs) = take (n - 1) xs" by (simp add: butlast_conv_take min.absorb1 min.absorb2) lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)" by (simp add: butlast_conv_take drop_take ac_simps) lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs" by (simp add: butlast_conv_take min.absorb1) lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)" by (simp add: butlast_conv_take drop_take ac_simps) lemma hd_drop_conv_nth: "n < length xs \ hd(drop n xs) = xs!n" by(simp add: hd_conv_nth) lemma set_take_subset_set_take: "m \ n \ set(take m xs) \ set(take n xs)" proof (induct xs arbitrary: m n) case (Cons x xs m n) then show ?case by (cases n) (auto simp: take_Cons) qed simp lemma set_take_subset: "set(take n xs) \ set xs" by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split) lemma set_drop_subset: "set(drop n xs) \ set xs" by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split) lemma set_drop_subset_set_drop: "m \ n \ set(drop m xs) \ set(drop n xs)" proof (induct xs arbitrary: m n) case (Cons x xs m n) then show ?case by (clarsimp simp: drop_Cons split: nat.split) (metis set_drop_subset subset_iff) qed simp lemma in_set_takeD: "x \ set(take n xs) \ x \ set xs" using set_take_subset by fast lemma in_set_dropD: "x \ set(drop n xs) \ x \ set xs" using set_drop_subset by fast lemma append_eq_conv_conj: "(xs @ ys = zs) = (xs = take (length xs) zs \ ys = drop (length xs) zs)" proof (induct xs arbitrary: zs) case (Cons x xs zs) then show ?case by (cases zs, auto) qed auto +lemma map_eq_append_conv: + "map f xs = ys @ zs \ (\us vs. xs = us @ vs \ ys = map f us \ zs = map f vs)" +proof - + have "map f xs \ ys @ zs \ map f xs \ ys @ zs \ map f xs \ ys @ zs \ map f xs = ys @ zs \ + (\bs bsa. xs = bs @ bsa \ ys = map f bs \ zs = map f bsa)" + by (metis append_eq_conv_conj append_take_drop_id drop_map take_map) + then show ?thesis + using map_append by blast +qed + +lemma append_eq_map_conv: + "ys @ zs = map f xs \ (\us vs. xs = us @ vs \ ys = map f us \ zs = map f vs)" +by (metis map_eq_append_conv) + lemma take_add: "take (i+j) xs = take i xs @ take j (drop i xs)" proof (induct xs arbitrary: i) case (Cons x xs i) then show ?case by (cases i, auto) qed auto lemma append_eq_append_conv_if: "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) = (if size xs\<^sub>1 \ size ys\<^sub>1 then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \ xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2 else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \ drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)" proof (induct xs\<^sub>1 arbitrary: ys\<^sub>1) case (Cons a xs\<^sub>1 ys\<^sub>1) then show ?case by (cases ys\<^sub>1, auto) qed auto lemma take_hd_drop: "n < length xs \ take n xs @ [hd (drop n xs)] = take (Suc n) xs" by (induct xs arbitrary: n) (simp_all add:drop_Cons split:nat.split) lemma id_take_nth_drop: "i < length xs \ xs = take i xs @ xs!i # drop (Suc i) xs" proof - assume si: "i < length xs" hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto moreover from si have "take (Suc i) xs = take i xs @ [xs!i]" using take_Suc_conv_app_nth by blast ultimately show ?thesis by auto qed lemma take_update_cancel[simp]: "n \ m \ take n (xs[m := y]) = take n xs" by(simp add: list_eq_iff_nth_eq) lemma drop_update_cancel[simp]: "n < m \ drop m (xs[n := x]) = drop m xs" by(simp add: list_eq_iff_nth_eq) lemma upd_conv_take_nth_drop: "i < length xs \ xs[i:=a] = take i xs @ a # drop (Suc i) xs" proof - assume i: "i < length xs" have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]" by(rule arg_cong[OF id_take_nth_drop[OF i]]) also have "\ = take i xs @ a # drop (Suc i) xs" using i by (simp add: list_update_append) finally show ?thesis . qed lemma take_update_swap: "take m (xs[n := x]) = (take m xs)[n := x]" proof (cases "n \ length xs") case False then show ?thesis by (simp add: upd_conv_take_nth_drop take_Cons drop_take min_def diff_Suc split: nat.split) qed auto -lemma drop_update_swap: +lemma drop_update_swap: assumes "m \ n" shows "drop m (xs[n := x]) = (drop m xs)[n-m := x]" proof (cases "n \ length xs") case False with assms show ?thesis by (simp add: upd_conv_take_nth_drop drop_take) qed auto lemma nth_image: "l \ size xs \ nth xs ` {0..\<^const>\takeWhile\ and \<^const>\dropWhile\\ lemma length_takeWhile_le: "length (takeWhile P xs) \ length xs" by (induct xs) auto lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs" by (induct xs) auto lemma takeWhile_append1 [simp]: "\x \ set xs; \P(x)\ \ takeWhile P (xs @ ys) = takeWhile P xs" by (induct xs) auto lemma takeWhile_append2 [simp]: "(\x. x \ set xs \ P x) \ takeWhile P (xs @ ys) = xs @ takeWhile P ys" by (induct xs) auto lemma takeWhile_tail: "\ P x \ takeWhile P (xs @ (x#l)) = takeWhile P xs" by (induct xs) auto lemma takeWhile_nth: "j < length (takeWhile P xs) \ takeWhile P xs ! j = xs ! j" by (metis nth_append takeWhile_dropWhile_id) lemma dropWhile_nth: "j < length (dropWhile P xs) \ dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))" by (metis add.commute nth_append_length_plus takeWhile_dropWhile_id) lemma length_dropWhile_le: "length (dropWhile P xs) \ length xs" by (induct xs) auto lemma dropWhile_append1 [simp]: "\x \ set xs; \P(x)\ \ dropWhile P (xs @ ys) = (dropWhile P xs)@ys" by (induct xs) auto lemma dropWhile_append2 [simp]: "(\x. x \ set xs \ P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys" by (induct xs) auto lemma dropWhile_append3: "\ P y \dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys" by (induct xs) auto lemma dropWhile_last: "x \ set xs \ \ P x \ last (dropWhile P xs) = last xs" by (auto simp add: dropWhile_append3 in_set_conv_decomp) lemma set_dropWhileD: "x \ set (dropWhile P xs) \ x \ set xs" by (induct xs) (auto split: if_split_asm) lemma set_takeWhileD: "x \ set (takeWhile P xs) \ x \ set xs \ P x" by (induct xs) (auto split: if_split_asm) lemma takeWhile_eq_all_conv[simp]: "(takeWhile P xs = xs) = (\x \ set xs. P x)" by(induct xs, auto) lemma dropWhile_eq_Nil_conv[simp]: "(dropWhile P xs = []) = (\x \ set xs. P x)" by(induct xs, auto) lemma dropWhile_eq_Cons_conv: "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \ \ P y)" by(induct xs, auto) lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)" by (induct xs) (auto dest: set_takeWhileD) lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)" by (induct xs) auto lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \ f) xs)" by (induct xs) auto lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \ f) xs)" by (induct xs) auto lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs" by (induct xs) auto lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs" by (induct xs) auto lemma hd_dropWhile: "dropWhile P xs \ [] \ \ P (hd (dropWhile P xs))" by (induct xs) auto lemma takeWhile_eq_filter: assumes "\ x. x \ set (dropWhile P xs) \ \ P x" shows "takeWhile P xs = filter P xs" proof - have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)" by simp have B: "filter P (dropWhile P xs) = []" unfolding filter_empty_conv using assms by blast have "filter P xs = takeWhile P xs" unfolding A filter_append B by (auto simp add: filter_id_conv dest: set_takeWhileD) thus ?thesis .. qed lemma takeWhile_eq_take_P_nth: "\ \ i. \ i < n ; i < length xs \ \ P (xs ! i) ; n < length xs \ \ P (xs ! n) \ \ takeWhile P xs = take n xs" proof (induct xs arbitrary: n) case Nil thus ?case by simp next case (Cons x xs) show ?case proof (cases n) case 0 with Cons show ?thesis by simp next case [simp]: (Suc n') have "P x" using Cons.prems(1)[of 0] by simp moreover have "takeWhile P xs = take n' xs" proof (rule Cons.hyps) fix i assume "i < n'" "i < length xs" thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp next assume "n' < length xs" thus "\ P (xs ! n')" using Cons by auto qed ultimately show ?thesis by simp qed qed lemma nth_length_takeWhile: "length (takeWhile P xs) < length xs \ \ P (xs ! length (takeWhile P xs))" by (induct xs) auto lemma length_takeWhile_less_P_nth: assumes all: "\ i. i < j \ P (xs ! i)" and "j \ length xs" shows "j \ length (takeWhile P xs)" proof (rule classical) assume "\ ?thesis" hence "length (takeWhile P xs) < length xs" using assms by simp thus ?thesis using all \\ ?thesis\ nth_length_takeWhile[of P xs] by auto qed lemma takeWhile_neq_rev: "\distinct xs; x \ set xs\ \ takeWhile (\y. y \ x) (rev xs) = rev (tl (dropWhile (\y. y \ x) xs))" by(induct xs) (auto simp: takeWhile_tail[where l="[]"]) lemma dropWhile_neq_rev: "\distinct xs; x \ set xs\ \ dropWhile (\y. y \ x) (rev xs) = x # rev (takeWhile (\y. y \ x) xs)" proof (induct xs) case (Cons a xs) then show ?case by(auto, subst dropWhile_append2, auto) qed simp lemma takeWhile_not_last: "distinct xs \ takeWhile (\y. y \ last xs) xs = butlast xs" by(induction xs rule: induct_list012) auto lemma takeWhile_cong [fundef_cong]: "\l = k; \x. x \ set l \ P x = Q x\ \ takeWhile P l = takeWhile Q k" by (induct k arbitrary: l) (simp_all) lemma dropWhile_cong [fundef_cong]: "\l = k; \x. x \ set l \ P x = Q x\ \ dropWhile P l = dropWhile Q k" by (induct k arbitrary: l, simp_all) lemma takeWhile_idem [simp]: "takeWhile P (takeWhile P xs) = takeWhile P xs" by (induct xs) auto lemma dropWhile_idem [simp]: "dropWhile P (dropWhile P xs) = dropWhile P xs" by (induct xs) auto subsubsection \\<^const>\zip\\ lemma zip_Nil [simp]: "zip [] ys = []" by (induct ys) auto lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys" by simp declare zip_Cons [simp del] lemma [code]: "zip [] ys = []" "zip xs [] = []" "zip (x # xs) (y # ys) = (x, y) # zip xs ys" by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+ lemma zip_Cons1: "zip (x#xs) ys = (case ys of [] \ [] | y#ys \ (x,y)#zip xs ys)" by(auto split:list.split) lemma length_zip [simp]: "length (zip xs ys) = min (length xs) (length ys)" by (induct xs ys rule:list_induct2') auto lemma zip_obtain_same_length: assumes "\zs ws n. length zs = length ws \ n = min (length xs) (length ys) \ zs = take n xs \ ws = take n ys \ P (zip zs ws)" shows "P (zip xs ys)" proof - let ?n = "min (length xs) (length ys)" have "P (zip (take ?n xs) (take ?n ys))" by (rule assms) simp_all moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases ys) simp_all qed ultimately show ?thesis by simp qed lemma zip_append1: "zip (xs @ ys) zs = zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)" by (induct xs zs rule:list_induct2') auto lemma zip_append2: "zip xs (ys @ zs) = zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs" by (induct xs ys rule:list_induct2') auto lemma zip_append [simp]: "[| length xs = length us |] ==> zip (xs@ys) (us@vs) = zip xs us @ zip ys vs" by (simp add: zip_append1) lemma zip_rev: "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)" by (induct rule:list_induct2, simp_all) lemma zip_map_map: "zip (map f xs) (map g ys) = map (\ (x, y). (f x, g y)) (zip xs ys)" proof (induct xs arbitrary: ys) case (Cons x xs) note Cons_x_xs = Cons.hyps show ?case proof (cases ys) case (Cons y ys') show ?thesis unfolding Cons using Cons_x_xs by simp qed simp qed simp lemma zip_map1: "zip (map f xs) ys = map (\(x, y). (f x, y)) (zip xs ys)" using zip_map_map[of f xs "\x. x" ys] by simp lemma zip_map2: "zip xs (map f ys) = map (\(x, y). (x, f y)) (zip xs ys)" using zip_map_map[of "\x. x" xs f ys] by simp lemma map_zip_map: "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)" by (auto simp: zip_map1) lemma map_zip_map2: "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)" by (auto simp: zip_map2) text\Courtesy of Andreas Lochbihler:\ lemma zip_same_conv_map: "zip xs xs = map (\x. (x, x)) xs" by(induct xs) auto lemma nth_zip [simp]: "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)" proof (induct ys arbitrary: i xs) case (Cons y ys) then show ?case by (cases xs) (simp_all add: nth.simps split: nat.split) qed auto lemma set_zip: "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}" by(simp add: set_conv_nth cong: rev_conj_cong) lemma zip_same: "((a,b) \ set (zip xs xs)) = (a \ set xs \ a = b)" by(induct xs) auto lemma zip_update: "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]" by (simp add: update_zip) lemma zip_replicate [simp]: "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)" proof (induct i arbitrary: j) case (Suc i) then show ?case by (cases j, auto) qed auto lemma zip_replicate1: "zip (replicate n x) ys = map (Pair x) (take n ys)" by(induction ys arbitrary: n)(case_tac [2] n, simp_all) lemma take_zip: "take n (zip xs ys) = zip (take n xs) (take n ys)" proof (induct n arbitrary: xs ys) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases ys) simp_all qed lemma drop_zip: "drop n (zip xs ys) = zip (drop n xs) (drop n ys)" proof (induct n arbitrary: xs ys) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases ys) simp_all qed lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \ fst) (zip xs ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \ snd) (zip xs ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma set_zip_leftD: "(x,y)\ set (zip xs ys) \ x \ set xs" by (induct xs ys rule:list_induct2') auto lemma set_zip_rightD: "(x,y)\ set (zip xs ys) \ y \ set ys" by (induct xs ys rule:list_induct2') auto lemma in_set_zipE: "(x,y) \ set(zip xs ys) \ (\ x \ set xs; y \ set ys \ \ R) \ R" by(blast dest: set_zip_leftD set_zip_rightD) lemma zip_map_fst_snd: "zip (map fst zs) (map snd zs) = zs" by (induct zs) simp_all lemma zip_eq_conv: "length xs = length ys \ zip xs ys = zs \ map fst zs = xs \ map snd zs = ys" by (auto simp add: zip_map_fst_snd) lemma in_set_zip: "p \ set (zip xs ys) \ (\n. xs ! n = fst p \ ys ! n = snd p \ n < length xs \ n < length ys)" by (cases p) (auto simp add: set_zip) lemma in_set_impl_in_set_zip1: assumes "length xs = length ys" assumes "x \ set xs" obtains y where "(x, y) \ set (zip xs ys)" proof - from assms have "x \ set (map fst (zip xs ys))" by simp from this that show ?thesis by fastforce qed lemma in_set_impl_in_set_zip2: assumes "length xs = length ys" assumes "y \ set ys" obtains x where "(x, y) \ set (zip xs ys)" proof - from assms have "y \ set (map snd (zip xs ys))" by simp from this that show ?thesis by fastforce qed lemma zip_eq_Nil_iff: "zip xs ys = [] \ xs = [] \ ys = []" by (cases xs; cases ys) simp_all lemma zip_eq_ConsE: assumes "zip xs ys = xy # xys" obtains x xs' y ys' where "xs = x # xs'" and "ys = y # ys'" and "xy = (x, y)" and "xys = zip xs' ys'" proof - from assms have "xs \ []" and "ys \ []" using zip_eq_Nil_iff [of xs ys] by simp_all then obtain x xs' y ys' where xs: "xs = x # xs'" and ys: "ys = y # ys'" by (cases xs; cases ys) auto with assms have "xy = (x, y)" and "xys = zip xs' ys'" by simp_all with xs ys show ?thesis .. qed lemma semilattice_map2: "semilattice (map2 (\<^bold>*))" if "semilattice (\<^bold>*)" for f (infixl "\<^bold>*" 70) proof - from that interpret semilattice f . show ?thesis proof show "map2 (\<^bold>*) (map2 (\<^bold>*) xs ys) zs = map2 (\<^bold>*) xs (map2 (\<^bold>*) ys zs)" for xs ys zs :: "'a list" proof (induction "zip xs (zip ys zs)" arbitrary: xs ys zs) case Nil from Nil [symmetric] show ?case by (auto simp add: zip_eq_Nil_iff) next case (Cons xyz xyzs) from Cons.hyps(2) [symmetric] show ?case by (rule zip_eq_ConsE) (erule zip_eq_ConsE, auto intro: Cons.hyps(1) simp add: ac_simps) qed show "map2 (\<^bold>*) xs ys = map2 (\<^bold>*) ys xs" for xs ys :: "'a list" proof (induction "zip xs ys" arbitrary: xs ys) case Nil then show ?case by (auto simp add: zip_eq_Nil_iff dest: sym) next case (Cons xy xys) from Cons.hyps(2) [symmetric] show ?case by (rule zip_eq_ConsE) (auto intro: Cons.hyps(1) simp add: ac_simps) qed show "map2 (\<^bold>*) xs xs = xs" for xs :: "'a list" by (induction xs) simp_all qed qed lemma pair_list_eqI: assumes "map fst xs = map fst ys" and "map snd xs = map snd ys" shows "xs = ys" proof - from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq) from this assms show ?thesis by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI) qed subsubsection \\<^const>\list_all2\\ lemma list_all2_lengthD [intro?]: "list_all2 P xs ys ==> length xs = length ys" by (simp add: list_all2_iff) lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])" by (simp add: list_all2_iff) lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])" by (simp add: list_all2_iff) lemma list_all2_Cons [iff, code]: "list_all2 P (x # xs) (y # ys) = (P x y \ list_all2 P xs ys)" by (auto simp add: list_all2_iff) lemma list_all2_Cons1: "list_all2 P (x # xs) ys = (\z zs. ys = z # zs \ P x z \ list_all2 P xs zs)" by (cases ys) auto lemma list_all2_Cons2: "list_all2 P xs (y # ys) = (\z zs. xs = z # zs \ P z y \ list_all2 P zs ys)" by (cases xs) auto lemma list_all2_induct [consumes 1, case_names Nil Cons, induct set: list_all2]: assumes P: "list_all2 P xs ys" assumes Nil: "R [] []" assumes Cons: "\x xs y ys. \P x y; list_all2 P xs ys; R xs ys\ \ R (x # xs) (y # ys)" shows "R xs ys" using P by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons) lemma list_all2_rev [iff]: "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys" by (simp add: list_all2_iff zip_rev cong: conj_cong) lemma list_all2_rev1: "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)" by (subst list_all2_rev [symmetric]) simp lemma list_all2_append1: "list_all2 P (xs @ ys) zs = (\us vs. zs = us @ vs \ length us = length xs \ length vs = length ys \ list_all2 P xs us \ list_all2 P ys vs)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (rule_tac x = "take (length xs) zs" in exI) apply (rule_tac x = "drop (length xs) zs" in exI) apply (force split: nat_diff_split simp add: list_all2_iff zip_append1) done next assume ?rhs then show ?lhs by (auto simp add: list_all2_iff) qed lemma list_all2_append2: "list_all2 P xs (ys @ zs) = (\us vs. xs = us @ vs \ length us = length ys \ length vs = length zs \ list_all2 P us ys \ list_all2 P vs zs)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (rule_tac x = "take (length ys) xs" in exI) apply (rule_tac x = "drop (length ys) xs" in exI) apply (force split: nat_diff_split simp add: list_all2_iff zip_append2) done next assume ?rhs then show ?lhs by (auto simp add: list_all2_iff) qed lemma list_all2_append: "length xs = length ys \ list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \ list_all2 P us vs)" by (induct rule:list_induct2, simp_all) lemma list_all2_appendI [intro?, trans]: "\ list_all2 P a b; list_all2 P c d \ \ list_all2 P (a@c) (b@d)" by (simp add: list_all2_append list_all2_lengthD) lemma list_all2_conv_all_nth: "list_all2 P xs ys = (length xs = length ys \ (\i < length xs. P (xs!i) (ys!i)))" by (force simp add: list_all2_iff set_zip) lemma list_all2_trans: assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c" shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs" (is "!!bs cs. PROP ?Q as bs cs") proof (induct as) fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs" show "!!cs. PROP ?Q (x # xs) bs cs" proof (induct bs) fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs" show "PROP ?Q (x # xs) (y # ys) cs" by (induct cs) (auto intro: tr I1 I2) qed simp qed simp lemma list_all2_all_nthI [intro?]: "length a = length b \ (\n. n < length a \ P (a!n) (b!n)) \ list_all2 P a b" by (simp add: list_all2_conv_all_nth) lemma list_all2I: "\x \ set (zip a b). case_prod P x \ length a = length b \ list_all2 P a b" by (simp add: list_all2_iff) lemma list_all2_nthD: "\ list_all2 P xs ys; p < size xs \ \ P (xs!p) (ys!p)" by (simp add: list_all2_conv_all_nth) lemma list_all2_nthD2: "\list_all2 P xs ys; p < size ys\ \ P (xs!p) (ys!p)" by (frule list_all2_lengthD) (auto intro: list_all2_nthD) lemma list_all2_map1: "list_all2 P (map f as) bs = list_all2 (\x y. P (f x) y) as bs" by (simp add: list_all2_conv_all_nth) lemma list_all2_map2: "list_all2 P as (map f bs) = list_all2 (\x y. P x (f y)) as bs" by (auto simp add: list_all2_conv_all_nth) lemma list_all2_refl [intro?]: "(\x. P x x) \ list_all2 P xs xs" by (simp add: list_all2_conv_all_nth) lemma list_all2_update_cong: "\ list_all2 P xs ys; P x y \ \ list_all2 P (xs[i:=x]) (ys[i:=y])" by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update) lemma list_all2_takeI [simp,intro?]: "list_all2 P xs ys \ list_all2 P (take n xs) (take n ys)" proof (induct xs arbitrary: n ys) case (Cons x xs) then show ?case by (cases n) (auto simp: list_all2_Cons1) qed auto lemma list_all2_dropI [simp,intro?]: "list_all2 P xs ys \ list_all2 P (drop n xs) (drop n ys)" proof (induct xs arbitrary: n ys) case (Cons x xs) then show ?case by (cases n) (auto simp: list_all2_Cons1) qed auto lemma list_all2_mono [intro?]: "list_all2 P xs ys \ (\xs ys. P xs ys \ Q xs ys) \ list_all2 Q xs ys" by (rule list.rel_mono_strong) lemma list_all2_eq: "xs = ys \ list_all2 (=) xs ys" by (induct xs ys rule: list_induct2') auto lemma list_eq_iff_zip_eq: "xs = ys \ length xs = length ys \ (\(x,y) \ set (zip xs ys). x = y)" by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong) lemma list_all2_same: "list_all2 P xs xs \ (\x\set xs. P x x)" by(auto simp add: list_all2_conv_all_nth set_conv_nth) lemma zip_assoc: "zip xs (zip ys zs) = map (\((x, y), z). (x, y, z)) (zip (zip xs ys) zs)" by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_commute: "zip xs ys = map (\(x, y). (y, x)) (zip ys xs)" by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_left_commute: "zip xs (zip ys zs) = map (\(y, (x, z)). (x, y, z)) (zip ys (zip xs zs))" by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_replicate2: "zip xs (replicate n y) = map (\x. (x, y)) (take n xs)" by(subst zip_commute)(simp add: zip_replicate1) subsubsection \\<^const>\List.product\ and \<^const>\product_lists\\ lemma product_concat_map: "List.product xs ys = concat (map (\x. map (\y. (x,y)) ys) xs)" by(induction xs) (simp)+ lemma set_product[simp]: "set (List.product xs ys) = set xs \ set ys" by (induct xs) auto lemma length_product [simp]: "length (List.product xs ys) = length xs * length ys" by (induct xs) simp_all lemma product_nth: assumes "n < length xs * length ys" shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))" using assms proof (induct xs arbitrary: n) case Nil then show ?case by simp next case (Cons x xs n) then have "length ys > 0" by auto with Cons show ?case by (auto simp add: nth_append not_less le_mod_geq le_div_geq) qed lemma in_set_product_lists_length: "xs \ set (product_lists xss) \ length xs = length xss" by (induct xss arbitrary: xs) auto lemma product_lists_set: "set (product_lists xss) = {xs. list_all2 (\x ys. x \ set ys) xs xss}" (is "?L = Collect ?R") proof (intro equalityI subsetI, unfold mem_Collect_eq) fix xs assume "xs \ ?L" then have "length xs = length xss" by (rule in_set_product_lists_length) from this \xs \ ?L\ show "?R xs" by (induct xs xss rule: list_induct2) auto next fix xs assume "?R xs" then show "xs \ ?L" by induct auto qed subsubsection \\<^const>\fold\ with natural argument order\ lemma fold_simps [code]: \ \eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala\ "fold f [] s = s" "fold f (x # xs) s = fold f xs (f x s)" by simp_all lemma fold_remove1_split: "\ \x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x; x \ set xs \ \ fold f xs = fold f (remove1 x xs) \ f x" by (induct xs) (auto simp add: comp_assoc) lemma fold_cong [fundef_cong]: "a = b \ xs = ys \ (\x. x \ set xs \ f x = g x) \ fold f xs a = fold g ys b" by (induct ys arbitrary: a b xs) simp_all lemma fold_id: "(\x. x \ set xs \ f x = id) \ fold f xs = id" by (induct xs) simp_all lemma fold_commute: "(\x. x \ set xs \ h \ g x = f x \ h) \ h \ fold g xs = fold f xs \ h" by (induct xs) (simp_all add: fun_eq_iff) lemma fold_commute_apply: assumes "\x. x \ set xs \ h \ g x = f x \ h" shows "h (fold g xs s) = fold f xs (h s)" proof - from assms have "h \ fold g xs = fold f xs \ h" by (rule fold_commute) then show ?thesis by (simp add: fun_eq_iff) qed lemma fold_invariant: "\ \x. x \ set xs \ Q x; P s; \x s. Q x \ P s \ P (f x s) \ \ P (fold f xs s)" by (induct xs arbitrary: s) simp_all lemma fold_append [simp]: "fold f (xs @ ys) = fold f ys \ fold f xs" by (induct xs) simp_all lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g \ f) xs" by (induct xs) simp_all lemma fold_filter: "fold f (filter P xs) = fold (\x. if P x then f x else id) xs" by (induct xs) simp_all lemma fold_rev: "(\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y) \ fold f (rev xs) = fold f xs" by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff) lemma fold_Cons_rev: "fold Cons xs = append (rev xs)" by (induct xs) simp_all lemma rev_conv_fold [code]: "rev xs = fold Cons xs []" by (simp add: fold_Cons_rev) lemma fold_append_concat_rev: "fold append xss = append (concat (rev xss))" by (induct xss) simp_all text \\<^const>\Finite_Set.fold\ and \<^const>\fold\\ lemma (in comp_fun_commute) fold_set_fold_remdups: "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb) lemma (in comp_fun_idem) fold_set_fold: "Finite_Set.fold f y (set xs) = fold f xs y" by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm) lemma union_set_fold [code]: "set xs \ A = fold Set.insert xs A" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by (simp add: union_fold_insert fold_set_fold) qed lemma union_coset_filter [code]: "List.coset xs \ A = List.coset (List.filter (\x. x \ A) xs)" by auto lemma minus_set_fold [code]: "A - set xs = fold Set.remove xs A" proof - interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove) show ?thesis by (simp add: minus_fold_remove [of _ A] fold_set_fold) qed lemma minus_coset_filter [code]: "A - List.coset xs = set (List.filter (\x. x \ A) xs)" by auto lemma inter_set_filter [code]: "A \ set xs = set (List.filter (\x. x \ A) xs)" by auto lemma inter_coset_fold [code]: "A \ List.coset xs = fold Set.remove xs A" by (simp add: Diff_eq [symmetric] minus_set_fold) lemma (in semilattice_set) set_eq_fold [code]: "F (set (x # xs)) = fold f xs x" proof - interpret comp_fun_idem f by standard (simp_all add: fun_eq_iff left_commute) show ?thesis by (simp add: eq_fold fold_set_fold) qed lemma (in complete_lattice) Inf_set_fold: "Inf (set xs) = fold inf xs top" proof - interpret comp_fun_idem "inf :: 'a \ 'a \ 'a" by (fact comp_fun_idem_inf) show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute) qed declare Inf_set_fold [where 'a = "'a set", code] lemma (in complete_lattice) Sup_set_fold: "Sup (set xs) = fold sup xs bot" proof - interpret comp_fun_idem "sup :: 'a \ 'a \ 'a" by (fact comp_fun_idem_sup) show ?thesis by (simp add: Sup_fold_sup fold_set_fold sup_commute) qed declare Sup_set_fold [where 'a = "'a set", code] lemma (in complete_lattice) INF_set_fold: "\(f ` set xs) = fold (inf \ f) xs top" using Inf_set_fold [of "map f xs"] by (simp add: fold_map) lemma (in complete_lattice) SUP_set_fold: "\(f ` set xs) = fold (sup \ f) xs bot" using Sup_set_fold [of "map f xs"] by (simp add: fold_map) subsubsection \Fold variants: \<^const>\foldr\ and \<^const>\foldl\\ text \Correspondence\ lemma foldr_conv_fold [code_abbrev]: "foldr f xs = fold f (rev xs)" by (induct xs) simp_all lemma foldl_conv_fold: "foldl f s xs = fold (\x s. f s x) xs s" by (induct xs arbitrary: s) simp_all lemma foldr_conv_foldl: \ \The ``Third Duality Theorem'' in Bird \& Wadler:\ "foldr f xs a = foldl (\x y. f y x) a (rev xs)" by (simp add: foldr_conv_fold foldl_conv_fold) lemma foldl_conv_foldr: "foldl f a xs = foldr (\x y. f y x) (rev xs) a" by (simp add: foldr_conv_fold foldl_conv_fold) lemma foldr_fold: "(\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y) \ foldr f xs = fold f xs" unfolding foldr_conv_fold by (rule fold_rev) lemma foldr_cong [fundef_cong]: "a = b \ l = k \ (\a x. x \ set l \ f x a = g x a) \ foldr f l a = foldr g k b" by (auto simp add: foldr_conv_fold intro!: fold_cong) lemma foldl_cong [fundef_cong]: "a = b \ l = k \ (\a x. x \ set l \ f a x = g a x) \ foldl f a l = foldl g b k" by (auto simp add: foldl_conv_fold intro!: fold_cong) lemma foldr_append [simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)" by (simp add: foldr_conv_fold) lemma foldl_append [simp]: "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys" by (simp add: foldl_conv_fold) lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g \ f) xs a" by (simp add: foldr_conv_fold fold_map rev_map) lemma foldr_filter: "foldr f (filter P xs) = foldr (\x. if P x then f x else id) xs" by (simp add: foldr_conv_fold rev_filter fold_filter) lemma foldl_map [code_unfold]: "foldl g a (map f xs) = foldl (\a x. g a (f x)) a xs" by (simp add: foldl_conv_fold fold_map comp_def) lemma concat_conv_foldr [code]: "concat xss = foldr append xss []" by (simp add: fold_append_concat_rev foldr_conv_fold) subsubsection \\<^const>\upt\\ lemma upt_rec[code]: "[i.. \simp does not terminate!\ by (induct j) auto lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n lemma upt_conv_Nil [simp]: "j \ i ==> [i.. j \ i)" by(induct j)simp_all lemma upt_eq_Cons_conv: "([i.. i = x \ [i+1.. j ==> [i..<(Suc j)] = [i.. \Only needed if \upt_Suc\ is deleted from the simpset.\ by simp lemma upt_conv_Cons: "i < j ==> [i.. \no precondition\ "m # n # ns = [m.. n # ns = [Suc m.. [i.. \LOOPS as a simprule, since \j \ j\.\ by (induct k) auto lemma length_upt [simp]: "length [i.. [i.. hd[i.. last[i.. n ==> take m [i..i. i + n) [0.. (map f [m..n. n - Suc 0) [Suc m..i. f (Suc i)) [0 ..< n]" by (induct n arbitrary: f) auto lemma nth_take_lemma: "k \ length xs \ k \ length ys \ (\i. i < k \ xs!i = ys!i) \ take k xs = take k ys" proof (induct k arbitrary: xs ys) case (Suc k) then show ?case apply (simp add: less_Suc_eq_0_disj) by (simp add: Suc.prems(3) take_Suc_conv_app_nth) qed simp lemma nth_equalityI: "\length xs = length ys; \i. i < length xs \ xs!i = ys!i\ \ xs = ys" by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all lemma map_nth: "map (\i. xs ! i) [0.. (\x y. \P x y; Q y x\ \ x = y); list_all2 P xs ys; list_all2 Q ys xs \ \ xs = ys" by (simp add: list_all2_conv_all_nth nth_equalityI) lemma take_equalityI: "(\i. take i xs = take i ys) ==> xs = ys" \ \The famous take-lemma.\ by (metis length_take min.commute order_refl take_all) lemma take_Cons': "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)" by (cases n) simp_all lemma drop_Cons': "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)" by (cases n) simp_all lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))" by (cases n) simp_all lemma take_Cons_numeral [simp]: "take (numeral v) (x # xs) = x # take (numeral v - 1) xs" by (simp add: take_Cons') lemma drop_Cons_numeral [simp]: "drop (numeral v) (x # xs) = drop (numeral v - 1) xs" by (simp add: drop_Cons') lemma nth_Cons_numeral [simp]: "(x # xs) ! numeral v = xs ! (numeral v - 1)" by (simp add: nth_Cons') subsubsection \\upto\: interval-list on \<^typ>\int\\ function upto :: "int \ int \ int list" ("(1[_../_])") where "upto i j = (if i \ j then i # [i+1..j] else [])" by auto termination by(relation "measure(%(i::int,j). nat(j - i + 1))") auto declare upto.simps[simp del] lemmas upto_rec_numeral [simp] = upto.simps[of "numeral m" "numeral n"] upto.simps[of "numeral m" "- numeral n"] upto.simps[of "- numeral m" "numeral n"] upto.simps[of "- numeral m" "- numeral n"] for m n lemma upto_empty[simp]: "j < i \ [i..j] = []" by(simp add: upto.simps) lemma upto_single[simp]: "[i..i] = [i]" by(simp add: upto.simps) lemma upto_Nil[simp]: "[i..j] = [] \ j < i" by (simp add: upto.simps) lemma upto_Nil2[simp]: "[] = [i..j] \ j < i" by (simp add: upto.simps) lemma upto_rec1: "i \ j \ [i..j] = i#[i+1..j]" by(simp add: upto.simps) lemma upto_rec2: "i \ j \ [i..j] = [i..j - 1]@[j]" proof(induct "nat(j-i)" arbitrary: i j) case 0 thus ?case by(simp add: upto.simps) next case (Suc n) hence "n = nat (j - (i + 1))" "i < j" by linarith+ from this(2) Suc.hyps(1)[OF this(1)] Suc(2,3) upto_rec1 show ?case by simp qed lemma length_upto[simp]: "length [i..j] = nat(j - i + 1)" by(induction i j rule: upto.induct) (auto simp: upto.simps) lemma set_upto[simp]: "set[i..j] = {i..j}" proof(induct i j rule:upto.induct) case (1 i j) from this show ?case unfolding upto.simps[of i j] by auto qed lemma nth_upto[simp]: "i + int k \ j \ [i..j] ! k = i + int k" apply(induction i j arbitrary: k rule: upto.induct) apply(subst upto_rec1) apply(auto simp add: nth_Cons') done -lemma upto_split1: +lemma upto_split1: "i \ j \ j \ k \ [i..k] = [i..j-1] @ [j..k]" proof (induction j rule: int_ge_induct) case base thus ?case by (simp add: upto_rec1) next case step thus ?case using upto_rec1 upto_rec2 by simp qed -lemma upto_split2: +lemma upto_split2: "i \ j \ j \ k \ [i..k] = [i..j] @ [j+1..k]" using upto_rec1 upto_rec2 upto_split1 by auto lemma upto_split3: "\ i \ j; j \ k \ \ [i..k] = [i..j-1] @ j # [j+1..k]" using upto_rec1 upto_split1 by auto text\Tail recursive version for code generation:\ definition upto_aux :: "int \ int \ int list \ int list" where "upto_aux i j js = [i..j] @ js" lemma upto_aux_rec [code]: "upto_aux i j js = (if j\<^const>\distinct\ and \<^const>\remdups\ and \<^const>\remdups_adj\\ lemma distinct_tl: "distinct xs \ distinct (tl xs)" by (cases xs) simp_all lemma distinct_append [simp]: "distinct (xs @ ys) = (distinct xs \ distinct ys \ set xs \ set ys = {})" by (induct xs) auto lemma distinct_rev[simp]: "distinct(rev xs) = distinct xs" by(induct xs) auto lemma set_remdups [simp]: "set (remdups xs) = set xs" by (induct xs) (auto simp add: insert_absorb) lemma distinct_remdups [iff]: "distinct (remdups xs)" by (induct xs) auto lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs" by (induct xs, auto) lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \ distinct xs" by (metis distinct_remdups distinct_remdups_id) lemma finite_distinct_list: "finite A \ \xs. set xs = A \ distinct xs" by (metis distinct_remdups finite_list set_remdups) lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])" by (induct x, auto) lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])" by (induct x, auto) lemma length_remdups_leq[iff]: "length(remdups xs) \ length xs" by (induct xs) auto lemma length_remdups_eq[iff]: "(length (remdups xs) = length xs) = (remdups xs = xs)" proof (induct xs) case (Cons a xs) then show ?case by simp (metis Suc_n_not_le_n impossible_Cons length_remdups_leq) qed auto lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)" by (induct xs) auto lemma distinct_map: "distinct(map f xs) = (distinct xs \ inj_on f (set xs))" by (induct xs) auto lemma distinct_map_filter: "distinct (map f xs) \ distinct (map f (filter P xs))" by (induct xs) auto lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)" by (induct xs) auto lemma distinct_upt[simp]: "distinct[i.. distinct (take i xs)" proof (induct xs arbitrary: i) case (Cons a xs) then show ?case by (metis Cons.prems append_take_drop_id distinct_append) qed auto lemma distinct_drop[simp]: "distinct xs \ distinct (drop i xs)" proof (induct xs arbitrary: i) case (Cons a xs) then show ?case by (metis Cons.prems append_take_drop_id distinct_append) qed auto lemma distinct_list_update: assumes d: "distinct xs" and a: "a \ set xs - {xs!i}" shows "distinct (xs[i:=a])" proof (cases "i < length xs") case True with a have anot: "a \ set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}" by simp (metis in_set_dropD in_set_takeD) show ?thesis proof (cases "a = xs!i") case True with d show ?thesis by auto next case False then show ?thesis - using d anot \i < length xs\ + using d anot \i < length xs\ apply (simp add: upd_conv_take_nth_drop) by (metis disjoint_insert(1) distinct_append id_take_nth_drop set_simps(2)) qed next case False with d show ?thesis by auto qed lemma distinct_concat: "\ distinct xs; \ ys. ys \ set xs \ distinct ys; \ ys zs. \ ys \ set xs ; zs \ set xs ; ys \ zs \ \ set ys \ set zs = {} \ \ distinct (concat xs)" by (induct xs) auto text \It is best to avoid this indexed version of distinct, but sometimes it is useful.\ lemma distinct_conv_nth: "distinct xs = (\i < size xs. \j < size xs. i \ j \ xs!i \ xs!j)" proof (induct xs) case (Cons x xs) show ?case apply (auto simp add: Cons nth_Cons split: nat.split_asm) apply (metis Suc_less_eq2 in_set_conv_nth less_not_refl zero_less_Suc)+ done qed auto lemma nth_eq_iff_index_eq: "\ distinct xs; i < length xs; j < length xs \ \ (xs!i = xs!j) = (i = j)" by(auto simp: distinct_conv_nth) lemma distinct_Ex1: "distinct xs \ x \ set xs \ (\!i. i < length xs \ xs ! i = x)" by (auto simp: in_set_conv_nth nth_eq_iff_index_eq) lemma inj_on_nth: "distinct xs \ \i \ I. i < length xs \ inj_on (nth xs) I" by (rule inj_onI) (simp add: nth_eq_iff_index_eq) lemma bij_betw_nth: assumes "distinct xs" "A = {.. distinct xs; n < length xs \ \ set(xs[n := x]) = insert x (set xs - {xs!n})" by(auto simp: set_eq_iff in_set_conv_nth nth_list_update nth_eq_iff_index_eq) lemma distinct_swap[simp]: "\ i < size xs; j < size xs \ \ distinct(xs[i := xs!j, j := xs!i]) = distinct xs" apply (simp add: distinct_conv_nth nth_list_update) apply safe apply metis+ done lemma set_swap[simp]: "\ i < size xs; j < size xs \ \ set(xs[i := xs!j, j := xs!i]) = set xs" by(simp add: set_conv_nth nth_list_update) metis lemma distinct_card: "distinct xs ==> card (set xs) = size xs" by (induct xs) auto lemma card_distinct: "card (set xs) = size xs ==> distinct xs" proof (induct xs) case (Cons x xs) show ?case proof (cases "x \ set xs") case False with Cons show ?thesis by simp next case True with Cons.prems have "card (set xs) = Suc (length xs)" by (simp add: card_insert_if split: if_split_asm) moreover have "card (set xs) \ length xs" by (rule card_length) ultimately have False by simp thus ?thesis .. qed qed simp lemma distinct_length_filter: "distinct xs \ length (filter P xs) = card ({x. P x} Int set xs)" by (induct xs) (auto) lemma not_distinct_decomp: "\ distinct ws \ \xs ys zs y. ws = xs@[y]@ys@[y]@zs" proof (induct n == "length ws" arbitrary:ws) case (Suc n ws) then show ?case using length_Suc_conv [of ws n] apply (auto simp: eq_commute) apply (metis append_Nil in_set_conv_decomp_first) by (metis append_Cons) qed simp lemma not_distinct_conv_prefix: defines "dec as xs y ys \ y \ set xs \ distinct xs \ as = xs @ y # ys" shows "\distinct as \ (\xs y ys. dec as xs y ys)" (is "?L = ?R") proof assume "?L" then show "?R" proof (induct "length as" arbitrary: as rule: less_induct) case less obtain xs ys zs y where decomp: "as = (xs @ y # ys) @ y # zs" using not_distinct_decomp[OF less.prems] by auto show ?case proof (cases "distinct (xs @ y # ys)") case True with decomp have "dec as (xs @ y # ys) y zs" by (simp add: dec_def) then show ?thesis by blast next case False with less decomp obtain xs' y' ys' where "dec (xs @ y # ys) xs' y' ys'" by atomize_elim auto with decomp have "dec as xs' y' (ys' @ y # zs)" by (simp add: dec_def) then show ?thesis by blast qed qed qed (auto simp: dec_def) lemma distinct_product: "distinct xs \ distinct ys \ distinct (List.product xs ys)" by (induct xs) (auto intro: inj_onI simp add: distinct_map) lemma distinct_product_lists: assumes "\xs \ set xss. distinct xs" shows "distinct (product_lists xss)" using assms proof (induction xss) case (Cons xs xss) note * = this then show ?case proof (cases "product_lists xss") case Nil then show ?thesis by (induct xs) simp_all next case (Cons ps pss) with * show ?thesis by (auto intro!: inj_onI distinct_concat simp add: distinct_map) qed qed simp lemma length_remdups_concat: "length (remdups (concat xss)) = card (\xs\set xss. set xs)" by (simp add: distinct_card [symmetric]) +lemma remdups_append2: + "remdups (xs @ remdups ys) = remdups (xs @ ys)" +by(induction xs) auto + lemma length_remdups_card_conv: "length(remdups xs) = card(set xs)" proof - have xs: "concat[xs] = xs" by simp from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp qed lemma remdups_remdups: "remdups (remdups xs) = remdups xs" by (induct xs) simp_all lemma distinct_butlast: assumes "distinct xs" shows "distinct (butlast xs)" proof (cases "xs = []") case False from \xs \ []\ obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto with \distinct xs\ show ?thesis by simp qed (auto) lemma remdups_map_remdups: "remdups (map f (remdups xs)) = remdups (map f xs)" by (induct xs) simp_all lemma distinct_zipI1: assumes "distinct xs" shows "distinct (zip xs ys)" proof (rule zip_obtain_same_length) fix xs' :: "'a list" and ys' :: "'b list" and n assume "length xs' = length ys'" assume "xs' = take n xs" with assms have "distinct xs'" by simp with \length xs' = length ys'\ show "distinct (zip xs' ys')" by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE) qed lemma distinct_zipI2: assumes "distinct ys" shows "distinct (zip xs ys)" proof (rule zip_obtain_same_length) fix xs' :: "'b list" and ys' :: "'a list" and n assume "length xs' = length ys'" assume "ys' = take n ys" with assms have "distinct ys'" by simp with \length xs' = length ys'\ show "distinct (zip xs' ys')" by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE) qed lemma set_take_disj_set_drop_if_distinct: "distinct vs \ i \ j \ set (take i vs) \ set (drop j vs) = {}" by (auto simp: in_set_conv_nth distinct_conv_nth) (* The next two lemmas help Sledgehammer. *) lemma distinct_singleton: "distinct [x]" by simp lemma distinct_length_2_or_more: "distinct (a # b # xs) \ (a \ b \ distinct (a # xs) \ distinct (b # xs))" by force lemma remdups_adj_altdef: "(remdups_adj xs = ys) \ (\f::nat => nat. mono f \ f ` {0 ..< size xs} = {0 ..< size ys} \ (\i < size xs. xs!i = ys!(f i)) \ (\i. i + 1 < size xs \ (xs!i = xs!(i+1) \ f i = f(i+1))))" (is "?L \ (\f. ?p f xs ys)") proof assume ?L then show "\f. ?p f xs ys" proof (induct xs arbitrary: ys rule: remdups_adj.induct) case (1 ys) thus ?case by (intro exI[of _ id]) (auto simp: mono_def) next case (2 x ys) thus ?case by (intro exI[of _ id]) (auto simp: mono_def) next case (3 x1 x2 xs ys) let ?xs = "x1 # x2 # xs" let ?cond = "x1 = x2" define zs where "zs = remdups_adj (x2 # xs)" from 3(1-2)[of zs] obtain f where p: "?p f (x2 # xs) zs" unfolding zs_def by (cases ?cond) auto then have f0: "f 0 = 0" by (intro mono_image_least[where f=f]) blast+ from p have mono: "mono f" and f_xs_zs: "f ` {0.. []" unfolding zs_def by (induct xs) auto let ?Succ = "if ?cond then id else Suc" let ?x1 = "if ?cond then id else Cons x1" let ?f = "\ i. if i = 0 then 0 else ?Succ (f (i - 1))" have ys: "ys = ?x1 zs" unfolding ys by (cases ?cond, auto) have mono: "mono ?f" using \mono f\ unfolding mono_def by auto show ?case unfolding ys proof (intro exI[of _ ?f] conjI allI impI) show "mono ?f" by fact next fix i assume i: "i < length ?xs" with p show "?xs ! i = ?x1 zs ! (?f i)" using zs0 by auto next fix i assume i: "i + 1 < length ?xs" with p show "(?xs ! i = ?xs ! (i + 1)) = (?f i = ?f (i + 1))" by (cases i) (auto simp: f0) next have id: "{0 ..< length (?x1 zs)} = insert 0 (?Succ ` {0 ..< length zs})" using zsne by (cases ?cond, auto) { fix i assume "i < Suc (length xs)" hence "Suc i \ {0.. Collect ((<) 0)" by auto from imageI[OF this, of "\i. ?Succ (f (i - Suc 0))"] have "?Succ (f i) \ (\i. ?Succ (f (i - Suc 0))) ` ({0.. Collect ((<) 0))" by auto } then show "?f ` {0 ..< length ?xs} = {0 ..< length (?x1 zs)}" unfolding id f_xs_zs[symmetric] by auto qed qed next assume "\ f. ?p f xs ys" then show ?L proof (induct xs arbitrary: ys rule: remdups_adj.induct) case 1 then show ?case by auto next case (2 x) then obtain f where f_img: "f ` {0 ..< size [x]} = {0 ..< size ys}" and f_nth: "\i. i < size [x] \ [x]!i = ys!(f i)" by blast have "length ys = card (f ` {0 ..< size [x]})" using f_img by auto then have *: "length ys = 1" by auto then have "f 0 = 0" using f_img by auto with * show ?case using f_nth by (cases ys) auto next case (3 x1 x2 xs) from "3.prems" obtain f where f_mono: "mono f" and f_img: "f ` {0..i. i < length (x1 # x2 # xs) \ (x1 # x2 # xs) ! i = ys ! f i" "\i. i + 1 < length (x1 # x2 #xs) \ ((x1 # x2 # xs) ! i = (x1 # x2 # xs) ! (i + 1)) = (f i = f (i + 1))" by blast show ?case proof cases assume "x1 = x2" let ?f' = "f \ Suc" have "remdups_adj (x1 # xs) = ys" proof (intro "3.hyps" exI conjI impI allI) show "mono ?f'" using f_mono by (simp add: mono_iff_le_Suc) next have "?f' ` {0 ..< length (x1 # xs)} = f ` {Suc 0 ..< length (x1 # x2 # xs)}" apply safe apply fastforce subgoal for \ x by (cases x) auto done also have "\ = f ` {0 ..< length (x1 # x2 # xs)}" proof - have "f 0 = f (Suc 0)" using \x1 = x2\ f_nth[of 0] by simp then show ?thesis apply safe apply fastforce subgoal for \ x by (cases x) auto done qed also have "\ = {0 ..< length ys}" by fact finally show "?f' ` {0 ..< length (x1 # xs)} = {0 ..< length ys}" . qed (insert f_nth[of "Suc i" for i], auto simp: \x1 = x2\) then show ?thesis using \x1 = x2\ by simp next assume "x1 \ x2" have "2 \ length ys" proof - have "2 = card {f 0, f 1}" using \x1 \ x2\ f_nth[of 0] by auto also have "\ \ card (f ` {0..< length (x1 # x2 # xs)})" by (rule card_mono) auto finally show ?thesis using f_img by simp qed have "f 0 = 0" using f_mono f_img by (rule mono_image_least) simp have "f (Suc 0) = Suc 0" proof (rule ccontr) assume "f (Suc 0) \ Suc 0" then have "Suc 0 < f (Suc 0)" using f_nth[of 0] \x1 \ x2\ \f 0 = 0\ by auto then have "\i. Suc 0 < f (Suc i)" using f_mono by (meson Suc_le_mono le0 less_le_trans monoD) then have "Suc 0 \ f i" for i using \f 0 = 0\ by (cases i) fastforce+ then have "Suc 0 \ f ` {0 ..< length (x1 # x2 # xs)}" by auto then show False using f_img \2 \ length ys\ by auto qed obtain ys' where "ys = x1 # x2 # ys'" using \2 \ length ys\ f_nth[of 0] f_nth[of 1] apply (cases ys) apply (rename_tac [2] ys', case_tac [2] ys') apply (auto simp: \f 0 = 0\ \f (Suc 0) = Suc 0\) done define f' where "f' x = f (Suc x) - 1" for x { fix i have "Suc 0 \ f (Suc 0)" using f_nth[of 0] \x1 \ x2\ \f 0 = 0\ by auto also have "\ \ f (Suc i)" using f_mono by (rule monoD) arith finally have "Suc 0 \ f (Suc i)" . } note Suc0_le_f_Suc = this { fix i have "f (Suc i) = Suc (f' i)" using Suc0_le_f_Suc[of i] by (auto simp: f'_def) } note f_Suc = this have "remdups_adj (x2 # xs) = (x2 # ys')" proof (intro "3.hyps" exI conjI impI allI) show "mono f'" using Suc0_le_f_Suc f_mono by (auto simp: f'_def mono_iff_le_Suc le_diff_iff) next have "f' ` {0 ..< length (x2 # xs)} = (\x. f x - 1) ` {0 ..< length (x1 # x2 #xs)}" by (auto simp: f'_def \f 0 = 0\ \f (Suc 0) = Suc 0\ image_def Bex_def less_Suc_eq_0_disj) also have "\ = (\x. x - 1) ` f ` {0 ..< length (x1 # x2 #xs)}" by (auto simp: image_comp) also have "\ = (\x. x - 1) ` {0 ..< length ys}" by (simp only: f_img) also have "\ = {0 ..< length (x2 # ys')}" using \ys = _\ by (fastforce intro: rev_image_eqI) finally show "f' ` {0 ..< length (x2 # xs)} = {0 ..< length (x2 # ys')}" . qed (insert f_nth[of "Suc i" for i] \x1 \ x2\, auto simp add: f_Suc \ys = _\) then show ?case using \ys = _\ \x1 \ x2\ by simp qed qed qed lemma hd_remdups_adj[simp]: "hd (remdups_adj xs) = hd xs" by (induction xs rule: remdups_adj.induct) simp_all lemma remdups_adj_Cons: "remdups_adj (x # xs) = (case remdups_adj xs of [] \ [x] | y # xs \ if x = y then y # xs else x # y # xs)" by (induct xs arbitrary: x) (auto split: list.splits) lemma remdups_adj_append_two: "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_adjacent: "Suc i < length (remdups_adj xs) \ remdups_adj xs ! i \ remdups_adj xs ! Suc i" proof (induction xs arbitrary: i rule: remdups_adj.induct) case (3 x y xs i) thus ?case by (cases i, cases "x = y") (simp, auto simp: hd_conv_nth[symmetric]) qed simp_all lemma remdups_adj_rev[simp]: "remdups_adj (rev xs) = rev (remdups_adj xs)" by (induct xs rule: remdups_adj.induct, simp_all add: remdups_adj_append_two) lemma remdups_adj_length[simp]: "length (remdups_adj xs) \ length xs" by (induct xs rule: remdups_adj.induct, auto) lemma remdups_adj_length_ge1[simp]: "xs \ [] \ length (remdups_adj xs) \ Suc 0" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_Nil_iff[simp]: "remdups_adj xs = [] \ xs = []" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_set[simp]: "set (remdups_adj xs) = set xs" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_Cons_alt[simp]: "x # tl (remdups_adj (x # xs)) = remdups_adj (x # xs)" by (induct xs rule: remdups_adj.induct, auto) lemma remdups_adj_distinct: "distinct xs \ remdups_adj xs = xs" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_append: "remdups_adj (xs\<^sub>1 @ x # xs\<^sub>2) = remdups_adj (xs\<^sub>1 @ [x]) @ tl (remdups_adj (x # xs\<^sub>2))" by (induct xs\<^sub>1 rule: remdups_adj.induct, simp_all) lemma remdups_adj_singleton: "remdups_adj xs = [x] \ xs = replicate (length xs) x" by (induct xs rule: remdups_adj.induct, auto split: if_split_asm) lemma remdups_adj_map_injective: assumes "inj f" shows "remdups_adj (map f xs) = map f (remdups_adj xs)" by (induct xs rule: remdups_adj.induct) (auto simp add: injD[OF assms]) lemma remdups_adj_replicate: "remdups_adj (replicate n x) = (if n = 0 then [] else [x])" by (induction n) (auto simp: remdups_adj_Cons) lemma remdups_upt [simp]: "remdups [m.. n") case False then show ?thesis by simp next case True then obtain q where "n = m + q" by (auto simp add: le_iff_add) moreover have "remdups [m..\<^const>\insert\\ lemma in_set_insert [simp]: "x \ set xs \ List.insert x xs = xs" by (simp add: List.insert_def) lemma not_in_set_insert [simp]: "x \ set xs \ List.insert x xs = x # xs" by (simp add: List.insert_def) lemma insert_Nil [simp]: "List.insert x [] = [x]" by simp lemma set_insert [simp]: "set (List.insert x xs) = insert x (set xs)" by (auto simp add: List.insert_def) lemma distinct_insert [simp]: "distinct (List.insert x xs) = distinct xs" by (simp add: List.insert_def) lemma insert_remdups: "List.insert x (remdups xs) = remdups (List.insert x xs)" by (simp add: List.insert_def) subsubsection \\<^const>\List.union\\ text\This is all one should need to know about union:\ lemma set_union[simp]: "set (List.union xs ys) = set xs \ set ys" unfolding List.union_def by(induct xs arbitrary: ys) simp_all lemma distinct_union[simp]: "distinct(List.union xs ys) = distinct ys" unfolding List.union_def by(induct xs arbitrary: ys) simp_all subsubsection \\<^const>\List.find\\ lemma find_None_iff: "List.find P xs = None \ \ (\x. x \ set xs \ P x)" proof (induction xs) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (fastforce split: if_splits) qed lemma find_Some_iff: "List.find P xs = Some x \ (\i x = xs!i \ (\j P (xs!j)))" proof (induction xs) case Nil thus ?case by simp next case (Cons x xs) thus ?case apply(auto simp: nth_Cons' split: if_splits) using diff_Suc_1[unfolded One_nat_def] less_Suc_eq_0_disj by fastforce qed lemma find_cong[fundef_cong]: assumes "xs = ys" and "\x. x \ set ys \ P x = Q x" shows "List.find P xs = List.find Q ys" proof (cases "List.find P xs") case None thus ?thesis by (metis find_None_iff assms) next case (Some x) hence "List.find Q ys = Some x" using assms by (auto simp add: find_Some_iff) thus ?thesis using Some by auto qed lemma find_dropWhile: "List.find P xs = (case dropWhile (Not \ P) xs of [] \ None | x # _ \ Some x)" by (induct xs) simp_all subsubsection \\<^const>\count_list\\ lemma count_notin[simp]: "x \ set xs \ count_list xs x = 0" by (induction xs) auto lemma count_le_length: "count_list xs x \ length xs" by (induction xs) auto lemma sum_count_set: "set xs \ X \ finite X \ sum (count_list xs) X = length xs" proof (induction xs arbitrary: X) case (Cons x xs) then show ?case apply (auto simp: sum.If_cases sum.remove) by (metis (no_types) Cons.IH Cons.prems(2) diff_eq sum.remove) qed simp subsubsection \\<^const>\List.extract\\ lemma extract_None_iff: "List.extract P xs = None \ \ (\ x\set xs. P x)" by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits) (metis in_set_conv_decomp) lemma extract_SomeE: "List.extract P xs = Some (ys, y, zs) \ xs = ys @ y # zs \ P y \ \ (\ y \ set ys. P y)" by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits) lemma extract_Some_iff: "List.extract P xs = Some (ys, y, zs) \ xs = ys @ y # zs \ P y \ \ (\ y \ set ys. P y)" by(auto simp: extract_def dropWhile_eq_Cons_conv dest: set_takeWhileD split: list.splits) lemma extract_Nil_code[code]: "List.extract P [] = None" by(simp add: extract_def) lemma extract_Cons_code[code]: "List.extract P (x # xs) = (if P x then Some ([], x, xs) else (case List.extract P xs of None \ None | Some (ys, y, zs) \ Some (x#ys, y, zs)))" by(auto simp add: extract_def comp_def split: list.splits) (metis dropWhile_eq_Nil_conv list.distinct(1)) subsubsection \\<^const>\remove1\\ lemma remove1_append: "remove1 x (xs @ ys) = (if x \ set xs then remove1 x xs @ ys else xs @ remove1 x ys)" by (induct xs) auto lemma remove1_commute: "remove1 x (remove1 y zs) = remove1 y (remove1 x zs)" by (induct zs) auto lemma in_set_remove1[simp]: "a \ b \ a \ set(remove1 b xs) = (a \ set xs)" by (induct xs) auto lemma set_remove1_subset: "set(remove1 x xs) \ set xs" by (induct xs) auto lemma set_remove1_eq [simp]: "distinct xs \ set(remove1 x xs) = set xs - {x}" by (induct xs) auto lemma length_remove1: "length(remove1 x xs) = (if x \ set xs then length xs - 1 else length xs)" by (induct xs) (auto dest!:length_pos_if_in_set) lemma remove1_filter_not[simp]: "\ P x \ remove1 x (filter P xs) = filter P xs" by(induct xs) auto lemma filter_remove1: "filter Q (remove1 x xs) = remove1 x (filter Q xs)" by (induct xs) auto lemma notin_set_remove1[simp]: "x \ set xs \ x \ set(remove1 y xs)" by(insert set_remove1_subset) fast lemma distinct_remove1[simp]: "distinct xs \ distinct(remove1 x xs)" by (induct xs) simp_all lemma remove1_remdups: "distinct xs \ remove1 x (remdups xs) = remdups (remove1 x xs)" by (induct xs) simp_all lemma remove1_idem: "x \ set xs \ remove1 x xs = xs" by (induct xs) simp_all subsubsection \\<^const>\removeAll\\ lemma removeAll_filter_not_eq: "removeAll x = filter (\y. x \ y)" proof fix xs show "removeAll x xs = filter (\y. x \ y) xs" by (induct xs) auto qed lemma removeAll_append[simp]: "removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys" by (induct xs) auto lemma set_removeAll[simp]: "set(removeAll x xs) = set xs - {x}" by (induct xs) auto lemma removeAll_id[simp]: "x \ set xs \ removeAll x xs = xs" by (induct xs) auto (* Needs count:: 'a \ 'a list \ nat lemma length_removeAll: "length(removeAll x xs) = length xs - count x xs" *) lemma removeAll_filter_not[simp]: "\ P x \ removeAll x (filter P xs) = filter P xs" by(induct xs) auto lemma distinct_removeAll: "distinct xs \ distinct (removeAll x xs)" by (simp add: removeAll_filter_not_eq) lemma distinct_remove1_removeAll: "distinct xs \ remove1 x xs = removeAll x xs" by (induct xs) simp_all lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \ map f (removeAll x xs) = removeAll (f x) (map f xs)" by (induct xs) (simp_all add:inj_on_def) lemma map_removeAll_inj: "inj f \ map f (removeAll x xs) = removeAll (f x) (map f xs)" by (rule map_removeAll_inj_on, erule subset_inj_on, rule subset_UNIV) lemma length_removeAll_less_eq [simp]: "length (removeAll x xs) \ length xs" by (simp add: removeAll_filter_not_eq) lemma length_removeAll_less [termination_simp]: "x \ set xs \ length (removeAll x xs) < length xs" by (auto dest: length_filter_less simp add: removeAll_filter_not_eq) subsubsection \\<^const>\replicate\\ lemma length_replicate [simp]: "length (replicate n x) = n" by (induct n) auto lemma replicate_eqI: assumes "length xs = n" and "\y. y \ set xs \ y = x" shows "xs = replicate n x" using assms proof (induct xs arbitrary: n) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases n) simp_all qed lemma Ex_list_of_length: "\xs. length xs = n" by (rule exI[of _ "replicate n undefined"]) simp lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)" by (induct n) auto lemma map_replicate_const: "map (\ x. k) lst = replicate (length lst) k" by (induct lst) auto lemma replicate_app_Cons_same: "(replicate n x) @ (x # xs) = x # replicate n x @ xs" by (induct n) auto lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x" by (induct n) (auto simp: replicate_app_Cons_same) lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x" by (induct n) auto text\Courtesy of Matthias Daum:\ lemma append_replicate_commute: "replicate n x @ replicate k x = replicate k x @ replicate n x" by (metis add.commute replicate_add) text\Courtesy of Andreas Lochbihler:\ lemma filter_replicate: "filter P (replicate n x) = (if P x then replicate n x else [])" by(induct n) auto lemma hd_replicate [simp]: "n \ 0 ==> hd (replicate n x) = x" by (induct n) auto lemma tl_replicate [simp]: "tl (replicate n x) = replicate (n - 1) x" by (induct n) auto lemma last_replicate [simp]: "n \ 0 ==> last (replicate n x) = x" by (atomize (full), induct n) auto lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x" by (induct n arbitrary: i)(auto simp: nth_Cons split: nat.split) text\Courtesy of Matthias Daum (2 lemmas):\ lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x" proof (cases "k \ i") case True then show ?thesis by (simp add: min_def) next case False then have "replicate k x = replicate i x @ replicate (k - i) x" by (simp add: replicate_add [symmetric]) then show ?thesis by (simp add: min_def) qed lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x" proof (induct k arbitrary: i) case (Suc k) then show ?case by (simp add: drop_Cons') qed simp lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}" by (induct n) auto lemma set_replicate [simp]: "n \ 0 ==> set (replicate n x) = {x}" by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc) lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})" by auto lemma in_set_replicate[simp]: "(x \ set (replicate n y)) = (x = y \ n \ 0)" by (simp add: set_replicate_conv_if) lemma Ball_set_replicate[simp]: "(\x \ set(replicate n a). P x) = (P a \ n=0)" by(simp add: set_replicate_conv_if) lemma Bex_set_replicate[simp]: "(\x \ set(replicate n a). P x) = (P a \ n\0)" by(simp add: set_replicate_conv_if) lemma replicate_append_same: "replicate i x @ [x] = x # replicate i x" by (induct i) simp_all lemma map_replicate_trivial: "map (\i. x) [0.. n=0" by (induct n) auto lemma empty_replicate[simp]: "([] = replicate n x) \ n=0" by (induct n) auto lemma replicate_eq_replicate[simp]: "(replicate m x = replicate n y) \ (m=n \ (m\0 \ x=y))" proof (induct m arbitrary: n) case (Suc m n) then show ?case by (induct n) auto qed simp lemma replicate_length_filter: "replicate (length (filter (\y. x = y) xs)) x = filter (\y. x = y) xs" by (induct xs) auto lemma comm_append_are_replicate: "\ xs \ []; ys \ []; xs @ ys = ys @ xs \ \ \m n zs. concat (replicate m zs) = xs \ concat (replicate n zs) = ys" proof (induction "length (xs @ ys)" arbitrary: xs ys rule: less_induct) case less define xs' ys' where "xs' = (if (length xs \ length ys) then xs else ys)" and "ys' = (if (length xs \ length ys) then ys else xs)" then have prems': "length xs' \ length ys'" "xs' @ ys' = ys' @ xs'" and "xs' \ []" and len: "length (xs @ ys) = length (xs' @ ys')" using less by (auto intro: less.hyps) from prems' obtain ws where "ys' = xs' @ ws" by (auto simp: append_eq_append_conv2) have "\m n zs. concat (replicate m zs) = xs' \ concat (replicate n zs) = ys'" proof (cases "ws = []") case True then have "concat (replicate 1 xs') = xs'" and "concat (replicate 1 xs') = ys'" using \ys' = xs' @ ws\ by auto then show ?thesis by blast next case False from \ys' = xs' @ ws\ and \xs' @ ys' = ys' @ xs'\ have "xs' @ ws = ws @ xs'" by simp then have "\m n zs. concat (replicate m zs) = xs' \ concat (replicate n zs) = ws" using False and \xs' \ []\ and \ys' = xs' @ ws\ and len by (intro less.hyps) auto then obtain m n zs where *: "concat (replicate m zs) = xs'" and "concat (replicate n zs) = ws" by blast then have "concat (replicate (m + n) zs) = ys'" using \ys' = xs' @ ws\ by (simp add: replicate_add) with * show ?thesis by blast qed then show ?case using xs'_def ys'_def by meson qed lemma comm_append_is_replicate: fixes xs ys :: "'a list" assumes "xs \ []" "ys \ []" assumes "xs @ ys = ys @ xs" shows "\n zs. n > 1 \ concat (replicate n zs) = xs @ ys" proof - obtain m n zs where "concat (replicate m zs) = xs" and "concat (replicate n zs) = ys" using comm_append_are_replicate[of xs ys, OF assms] by blast then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys" using \xs \ []\ and \ys \ []\ by (auto simp: replicate_add) then show ?thesis by blast qed lemma Cons_replicate_eq: "x # xs = replicate n y \ x = y \ n > 0 \ xs = replicate (n - 1) x" by (induct n) auto lemma replicate_length_same: "(\y\set xs. y = x) \ replicate (length xs) x = xs" by (induct xs) simp_all lemma foldr_replicate [simp]: "foldr f (replicate n x) = f x ^^ n" by (induct n) (simp_all) lemma fold_replicate [simp]: "fold f (replicate n x) = f x ^^ n" by (subst foldr_fold [symmetric]) simp_all subsubsection \\<^const>\enumerate\\ lemma enumerate_simps [simp, code]: "enumerate n [] = []" "enumerate n (x # xs) = (n, x) # enumerate (Suc n) xs" by (simp_all add: enumerate_eq_zip upt_rec) lemma length_enumerate [simp]: "length (enumerate n xs) = length xs" by (simp add: enumerate_eq_zip) lemma map_fst_enumerate [simp]: "map fst (enumerate n xs) = [n.. set (enumerate n xs) \ n \ fst p \ fst p < length xs + n \ nth xs (fst p - n) = snd p" proof - { fix m assume "n \ m" moreover assume "m < length xs + n" ultimately have "[n.. xs ! (m - n) = xs ! (m - n) \ m - n < length xs" by auto then have "\q. [n.. xs ! q = xs ! (m - n) \ q < length xs" .. } then show ?thesis by (cases p) (auto simp add: enumerate_eq_zip in_set_zip) qed lemma nth_enumerate_eq: "m < length xs \ enumerate n xs ! m = (n + m, xs ! m)" by (simp add: enumerate_eq_zip) lemma enumerate_replicate_eq: "enumerate n (replicate m a) = map (\q. (q, a)) [n..k. (k, f k)) [n.. m") (simp_all add: zip_map2 zip_same_conv_map enumerate_eq_zip) subsubsection \\<^const>\rotate1\ and \<^const>\rotate\\ lemma rotate0[simp]: "rotate 0 = id" by(simp add:rotate_def) lemma rotate_Suc[simp]: "rotate (Suc n) xs = rotate1(rotate n xs)" by(simp add:rotate_def) lemma rotate_add: "rotate (m+n) = rotate m \ rotate n" by(simp add:rotate_def funpow_add) lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs" by(simp add:rotate_add) +lemma rotate1_map: "rotate1 (map f xs) = map f (rotate1 xs)" +by(cases xs) simp_all + lemma rotate1_rotate_swap: "rotate1 (rotate n xs) = rotate n (rotate1 xs)" by(simp add:rotate_def funpow_swap1) lemma rotate1_length01[simp]: "length xs \ 1 \ rotate1 xs = xs" by(cases xs) simp_all lemma rotate_length01[simp]: "length xs \ 1 \ rotate n xs = xs" by (induct n) (simp_all add:rotate_def) lemma rotate1_hd_tl: "xs \ [] \ rotate1 xs = tl xs @ [hd xs]" by (cases xs) simp_all lemma rotate_drop_take: "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs" proof (induct n) case (Suc n) show ?case proof (cases "xs = []") case False then show ?thesis proof (cases "n mod length xs = 0") case True then show ?thesis apply (simp add: mod_Suc) by (simp add: False Suc.hyps drop_Suc rotate1_hd_tl take_Suc) next case False with \xs \ []\ Suc show ?thesis by (simp add: rotate_def mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric] take_hd_drop linorder_not_le) qed qed simp qed simp lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs" by(simp add:rotate_drop_take) lemma rotate_id[simp]: "n mod length xs = 0 \ rotate n xs = xs" by(simp add:rotate_drop_take) lemma length_rotate1[simp]: "length(rotate1 xs) = length xs" by (cases xs) simp_all lemma length_rotate[simp]: "length(rotate n xs) = length xs" by (induct n arbitrary: xs) (simp_all add:rotate_def) lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs" by (cases xs) auto lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs" by (induct n) (simp_all add:rotate_def) lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)" by(simp add:rotate_drop_take take_map drop_map) lemma set_rotate1[simp]: "set(rotate1 xs) = set xs" by (cases xs) auto lemma set_rotate[simp]: "set(rotate n xs) = set xs" by (induct n) (simp_all add:rotate_def) lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])" by (cases xs) auto lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])" by (induct n) (simp_all add:rotate_def) lemma rotate_rev: "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)" proof (cases "length xs = 0 \ n mod length xs = 0") case False then show ?thesis by(simp add:rotate_drop_take rev_drop rev_take) qed force lemma hd_rotate_conv_nth: assumes "xs \ []" shows "hd(rotate n xs) = xs!(n mod length xs)" proof - have "n mod length xs < length xs" using assms by simp then show ?thesis by (metis drop_eq_Nil hd_append2 hd_drop_conv_nth leD rotate_drop_take) qed lemma rotate_append: "rotate (length l) (l @ q) = q @ l" by (induct l arbitrary: q) (auto simp add: rotate1_rotate_swap) subsubsection \\<^const>\nths\ --- a generalization of \<^const>\nth\ to sets\ lemma nths_empty [simp]: "nths xs {} = []" by (auto simp add: nths_def) lemma nths_nil [simp]: "nths [] A = []" by (auto simp add: nths_def) lemma nths_all: "\i < length xs. i \ I \ nths xs I = xs" apply (simp add: nths_def) apply (subst filter_True) apply (clarsimp simp: in_set_zip subset_iff) apply simp done lemma length_nths: "length (nths xs I) = card{i. i < length xs \ i \ I}" by(simp add: nths_def length_filter_conv_card cong:conj_cong) lemma nths_shift_lemma_Suc: "map fst (filter (\p. P(Suc(snd p))) (zip xs is)) = map fst (filter (\p. P(snd p)) (zip xs (map Suc is)))" proof (induct xs arbitrary: "is") case (Cons x xs "is") show ?case by (cases "is") (auto simp add: Cons.hyps) qed simp lemma nths_shift_lemma: "map fst (filter (\p. snd p \ A) (zip xs [i..p. snd p + i \ A) (zip xs [0.. A}" unfolding nths_def proof (induct l' rule: rev_induct) case (snoc x xs) then show ?case by (simp add: upt_add_eq_append[of 0] nths_shift_lemma add.commute) qed auto lemma nths_Cons: "nths (x # l) A = (if 0 \ A then [x] else []) @ nths l {j. Suc j \ A}" proof (induct l rule: rev_induct) case (snoc x xs) then show ?case by (simp flip: append_Cons add: nths_append) qed (auto simp: nths_def) lemma nths_map: "nths (map f xs) I = map f (nths xs I)" by(induction xs arbitrary: I) (simp_all add: nths_Cons) lemma set_nths: "set(nths xs I) = {xs!i|i. i i \ I}" by (induct xs arbitrary: I) (auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc) lemma set_nths_subset: "set(nths xs I) \ set xs" by(auto simp add:set_nths) lemma notin_set_nthsI[simp]: "x \ set xs \ x \ set(nths xs I)" by(auto simp add:set_nths) lemma in_set_nthsD: "x \ set(nths xs I) \ x \ set xs" by(auto simp add:set_nths) lemma nths_singleton [simp]: "nths [x] A = (if 0 \ A then [x] else [])" by (simp add: nths_Cons) lemma distinct_nthsI[simp]: "distinct xs \ distinct (nths xs I)" by (induct xs arbitrary: I) (auto simp: nths_Cons) lemma nths_upt_eq_take [simp]: "nths l {.. A. \j \ B. card {i' \ A. i' < i} = j}" apply(induction xs arbitrary: A B) apply(auto simp add: nths_Cons card_less_Suc card_less_Suc2) done lemma drop_eq_nths: "drop n xs = nths xs {i. i \ n}" apply(induction xs arbitrary: n) apply (auto simp add: nths_Cons nths_all drop_Cons' intro: arg_cong2[where f=nths, OF refl]) done lemma nths_drop: "nths (drop n xs) I = nths xs ((+) n ` I)" by(force simp: drop_eq_nths nths_nths simp flip: atLeastLessThan_iff intro: arg_cong2[where f=nths, OF refl]) lemma filter_eq_nths: "filter P xs = nths xs {i. i P(xs!i)}" by(induction xs) (auto simp: nths_Cons) lemma filter_in_nths: "distinct xs \ filter (%x. x \ set (nths xs s)) xs = nths xs s" proof (induct xs arbitrary: s) case Nil thus ?case by simp next case (Cons a xs) then have "\x. x \ set xs \ x \ a" by auto with Cons show ?case by(simp add: nths_Cons cong:filter_cong) qed subsubsection \\<^const>\subseqs\ and \<^const>\List.n_lists\\ lemma length_subseqs: "length (subseqs xs) = 2 ^ length xs" by (induct xs) (simp_all add: Let_def) lemma subseqs_powset: "set ` set (subseqs xs) = Pow (set xs)" proof - have aux: "\x A. set ` Cons x ` A = insert x ` set ` A" by (auto simp add: image_def) have "set (map set (subseqs xs)) = Pow (set xs)" by (induct xs) (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map) then show ?thesis by simp qed lemma distinct_set_subseqs: assumes "distinct xs" shows "distinct (map set (subseqs xs))" proof (rule card_distinct) have "finite (set xs)" .. then have "card (Pow (set xs)) = 2 ^ card (set xs)" by (rule card_Pow) with assms distinct_card [of xs] have "card (Pow (set xs)) = 2 ^ length xs" by simp then show "card (set (map set (subseqs xs))) = length (map set (subseqs xs))" by (simp add: subseqs_powset length_subseqs) qed lemma n_lists_Nil [simp]: "List.n_lists n [] = (if n = 0 then [[]] else [])" by (induct n) simp_all lemma length_n_lists_elem: "ys \ set (List.n_lists n xs) \ length ys = n" by (induct n arbitrary: ys) auto lemma set_n_lists: "set (List.n_lists n xs) = {ys. length ys = n \ set ys \ set xs}" proof (rule set_eqI) fix ys :: "'a list" show "ys \ set (List.n_lists n xs) \ ys \ {ys. length ys = n \ set ys \ set xs}" proof - have "ys \ set (List.n_lists n xs) \ length ys = n" by (induct n arbitrary: ys) auto moreover have "\x. ys \ set (List.n_lists n xs) \ x \ set ys \ x \ set xs" by (induct n arbitrary: ys) auto moreover have "set ys \ set xs \ ys \ set (List.n_lists (length ys) xs)" by (induct ys) auto ultimately show ?thesis by auto qed qed lemma subseqs_refl: "xs \ set (subseqs xs)" by (induct xs) (simp_all add: Let_def) lemma subset_subseqs: "X \ set xs \ X \ set ` set (subseqs xs)" unfolding subseqs_powset by simp lemma Cons_in_subseqsD: "y # ys \ set (subseqs xs) \ ys \ set (subseqs xs)" by (induct xs) (auto simp: Let_def) lemma subseqs_distinctD: "\ ys \ set (subseqs xs); distinct xs \ \ distinct ys" proof (induct xs arbitrary: ys) case (Cons x xs ys) then show ?case by (auto simp: Let_def) (metis Pow_iff contra_subsetD image_eqI subseqs_powset) qed simp subsubsection \\<^const>\splice\\ lemma splice_Nil2 [simp]: "splice xs [] = xs" by (cases xs) simp_all lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys" by (induct xs ys rule: splice.induct) auto lemma split_Nil_iff[simp]: "splice xs ys = [] \ xs = [] \ ys = []" by (induct xs ys rule: splice.induct) auto lemma splice_replicate[simp]: "splice (replicate m x) (replicate n x) = replicate (m+n) x" apply(induction "replicate m x" "replicate n x" arbitrary: m n rule: splice.induct) apply (auto simp add: Cons_replicate_eq dest: gr0_implies_Suc) done subsubsection \\<^const>\shuffles\\ lemma shuffles_commutes: "shuffles xs ys = shuffles ys xs" by (induction xs ys rule: shuffles.induct) (simp_all add: Un_commute) lemma Nil_in_shuffles[simp]: "[] \ shuffles xs ys \ xs = [] \ ys = []" by (induct xs ys rule: shuffles.induct) auto lemma shufflesE: "zs \ shuffles xs ys \ (zs = xs \ ys = [] \ P) \ (zs = ys \ xs = [] \ P) \ (\x xs' z zs'. xs = x # xs' \ zs = z # zs' \ x = z \ zs' \ shuffles xs' ys \ P) \ (\y ys' z zs'. ys = y # ys' \ zs = z # zs' \ y = z \ zs' \ shuffles xs ys' \ P) \ P" by (induct xs ys rule: shuffles.induct) auto lemma Cons_in_shuffles_iff: "z # zs \ shuffles xs ys \ (xs \ [] \ hd xs = z \ zs \ shuffles (tl xs) ys \ ys \ [] \ hd ys = z \ zs \ shuffles xs (tl ys))" by (induct xs ys rule: shuffles.induct) auto lemma splice_in_shuffles [simp, intro]: "splice xs ys \ shuffles xs ys" by (induction xs ys rule: splice.induct) (simp_all add: Cons_in_shuffles_iff shuffles_commutes) lemma Nil_in_shufflesI: "xs = [] \ ys = [] \ [] \ shuffles xs ys" by simp lemma Cons_in_shuffles_leftI: "zs \ shuffles xs ys \ z # zs \ shuffles (z # xs) ys" by (cases ys) auto lemma Cons_in_shuffles_rightI: "zs \ shuffles xs ys \ z # zs \ shuffles xs (z # ys)" by (cases xs) auto lemma finite_shuffles [simp, intro]: "finite (shuffles xs ys)" by (induction xs ys rule: shuffles.induct) simp_all lemma length_shuffles: "zs \ shuffles xs ys \ length zs = length xs + length ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma set_shuffles: "zs \ shuffles xs ys \ set zs = set xs \ set ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma distinct_disjoint_shuffles: assumes "distinct xs" "distinct ys" "set xs \ set ys = {}" "zs \ shuffles xs ys" shows "distinct zs" using assms proof (induction xs ys arbitrary: zs rule: shuffles.induct) case (3 x xs y ys) show ?case proof (cases zs) case (Cons z zs') with "3.prems" and "3.IH"[of zs'] show ?thesis by (force dest: set_shuffles) qed simp_all qed simp_all lemma Cons_shuffles_subset1: "(#) x ` shuffles xs ys \ shuffles (x # xs) ys" by (cases ys) auto lemma Cons_shuffles_subset2: "(#) y ` shuffles xs ys \ shuffles xs (y # ys)" by (cases xs) auto lemma filter_shuffles: "filter P ` shuffles xs ys = shuffles (filter P xs) (filter P ys)" proof - have *: "filter P ` (#) x ` A = (if P x then (#) x ` filter P ` A else filter P ` A)" for x A by (auto simp: image_image) show ?thesis by (induction xs ys rule: shuffles.induct) (simp_all split: if_splits add: image_Un * Un_absorb1 Un_absorb2 Cons_shuffles_subset1 Cons_shuffles_subset2) qed lemma filter_shuffles_disjoint1: assumes "set xs \ set ys = {}" "zs \ shuffles xs ys" shows "filter (\x. x \ set xs) zs = xs" (is "filter ?P _ = _") and "filter (\x. x \ set xs) zs = ys" (is "filter ?Q _ = _") using assms proof - from assms have "filter ?P zs \ filter ?P ` shuffles xs ys" by blast also have "filter ?P ` shuffles xs ys = shuffles (filter ?P xs) (filter ?P ys)" by (rule filter_shuffles) also have "filter ?P xs = xs" by (rule filter_True) simp_all also have "filter ?P ys = []" by (rule filter_False) (insert assms(1), auto) also have "shuffles xs [] = {xs}" by simp finally show "filter ?P zs = xs" by simp next from assms have "filter ?Q zs \ filter ?Q ` shuffles xs ys" by blast also have "filter ?Q ` shuffles xs ys = shuffles (filter ?Q xs) (filter ?Q ys)" by (rule filter_shuffles) also have "filter ?Q ys = ys" by (rule filter_True) (insert assms(1), auto) also have "filter ?Q xs = []" by (rule filter_False) (insert assms(1), auto) also have "shuffles [] ys = {ys}" by simp finally show "filter ?Q zs = ys" by simp qed lemma filter_shuffles_disjoint2: assumes "set xs \ set ys = {}" "zs \ shuffles xs ys" shows "filter (\x. x \ set ys) zs = ys" "filter (\x. x \ set ys) zs = xs" using filter_shuffles_disjoint1[of ys xs zs] assms by (simp_all add: shuffles_commutes Int_commute) lemma partition_in_shuffles: "xs \ shuffles (filter P xs) (filter (\x. \P x) xs)" proof (induction xs) case (Cons x xs) show ?case proof (cases "P x") case True hence "x # xs \ (#) x ` shuffles (filter P xs) (filter (\x. \P x) xs)" by (intro imageI Cons.IH) also have "\ \ shuffles (filter P (x # xs)) (filter (\x. \P x) (x # xs))" by (simp add: True Cons_shuffles_subset1) finally show ?thesis . next case False hence "x # xs \ (#) x ` shuffles (filter P xs) (filter (\x. \P x) xs)" by (intro imageI Cons.IH) also have "\ \ shuffles (filter P (x # xs)) (filter (\x. \P x) (x # xs))" by (simp add: False Cons_shuffles_subset2) finally show ?thesis . qed qed auto lemma inv_image_partition: assumes "\x. x \ set xs \ P x" "\y. y \ set ys \ \P y" shows "partition P -` {(xs, ys)} = shuffles xs ys" proof (intro equalityI subsetI) fix zs assume zs: "zs \ shuffles xs ys" hence [simp]: "set zs = set xs \ set ys" by (rule set_shuffles) from assms have "filter P zs = filter (\x. x \ set xs) zs" "filter (\x. \P x) zs = filter (\x. x \ set ys) zs" by (intro filter_cong refl; force)+ moreover from assms have "set xs \ set ys = {}" by auto ultimately show "zs \ partition P -` {(xs, ys)}" using zs by (simp add: o_def filter_shuffles_disjoint1 filter_shuffles_disjoint2) next fix zs assume "zs \ partition P -` {(xs, ys)}" thus "zs \ shuffles xs ys" using partition_in_shuffles[of zs] by (auto simp: o_def) qed subsubsection \Transpose\ function transpose where "transpose [] = []" | "transpose ([] # xss) = transpose xss" | "transpose ((x#xs) # xss) = (x # [h. (h#t) \ xss]) # transpose (xs # [t. (h#t) \ xss])" by pat_completeness auto lemma transpose_aux_filter_head: "concat (map (case_list [] (\h t. [h])) xss) = map (\xs. hd xs) (filter (\ys. ys \ []) xss)" by (induct xss) (auto split: list.split) lemma transpose_aux_filter_tail: "concat (map (case_list [] (\h t. [t])) xss) = map (\xs. tl xs) (filter (\ys. ys \ []) xss)" by (induct xss) (auto split: list.split) lemma transpose_aux_max: "max (Suc (length xs)) (foldr (\xs. max (length xs)) xss 0) = Suc (max (length xs) (foldr (\x. max (length x - Suc 0)) (filter (\ys. ys \ []) xss) 0))" (is "max _ ?foldB = Suc (max _ ?foldA)") proof (cases "(filter (\ys. ys \ []) xss) = []") case True hence "foldr (\xs. max (length xs)) xss 0 = 0" proof (induct xss) case (Cons x xs) then have "x = []" by (cases x) auto with Cons show ?case by auto qed simp thus ?thesis using True by simp next case False have foldA: "?foldA = foldr (\x. max (length x)) (filter (\ys. ys \ []) xss) 0 - 1" by (induct xss) auto have foldB: "?foldB = foldr (\x. max (length x)) (filter (\ys. ys \ []) xss) 0" by (induct xss) auto have "0 < ?foldB" proof - from False obtain z zs where zs: "(filter (\ys. ys \ []) xss) = z#zs" by (auto simp: neq_Nil_conv) hence "z \ set (filter (\ys. ys \ []) xss)" by auto hence "z \ []" by auto thus ?thesis unfolding foldB zs by (auto simp: max_def intro: less_le_trans) qed thus ?thesis unfolding foldA foldB max_Suc_Suc[symmetric] by simp qed termination transpose by (relation "measure (\xs. foldr (\xs. max (length xs)) xs 0 + length xs)") (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max less_Suc_eq_le) lemma transpose_empty: "(transpose xs = []) \ (\x \ set xs. x = [])" by (induct rule: transpose.induct) simp_all lemma length_transpose: fixes xs :: "'a list list" shows "length (transpose xs) = foldr (\xs. max (length xs)) xs 0" by (induct rule: transpose.induct) (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max max_Suc_Suc[symmetric] simp del: max_Suc_Suc) lemma nth_transpose: fixes xs :: "'a list list" assumes "i < length (transpose xs)" shows "transpose xs ! i = map (\xs. xs ! i) (filter (\ys. i < length ys) xs)" using assms proof (induct arbitrary: i rule: transpose.induct) case (3 x xs xss) define XS where "XS = (x # xs) # xss" hence [simp]: "XS \ []" by auto thus ?case proof (cases i) case 0 thus ?thesis by (simp add: transpose_aux_filter_head hd_conv_nth) next case (Suc j) have *: "\xss. xs # map tl xss = map tl ((x#xs)#xss)" by simp have **: "\xss. (x#xs) # filter (\ys. ys \ []) xss = filter (\ys. ys \ []) ((x#xs)#xss)" by simp { fix x have "Suc j < length x \ x \ [] \ j < length x - Suc 0" by (cases x) simp_all } note *** = this have j_less: "j < length (transpose (xs # concat (map (case_list [] (\h t. [t])) xss)))" using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc) show ?thesis unfolding transpose.simps \i = Suc j\ nth_Cons_Suc "3.hyps"[OF j_less] apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric]) by (simp add: nth_tl) qed qed simp_all lemma transpose_map_map: "transpose (map (map f) xs) = map (map f) (transpose xs)" proof (rule nth_equalityI) have [simp]: "length (transpose (map (map f) xs)) = length (transpose xs)" by (simp add: length_transpose foldr_map comp_def) show "length (transpose (map (map f) xs)) = length (map (map f) (transpose xs))" by simp fix i assume "i < length (transpose (map (map f) xs))" thus "transpose (map (map f) xs) ! i = map (map f) (transpose xs) ! i" by (simp add: nth_transpose filter_map comp_def) qed subsubsection \\<^const>\min\ and \<^const>\arg_min\\ lemma min_list_Min: "xs \ [] \ min_list xs = Min (set xs)" by (induction xs rule: induct_list012)(auto) lemma f_arg_min_list_f: "xs \ [] \ f (arg_min_list f xs) = Min (f ` (set xs))" by(induction f xs rule: arg_min_list.induct) (auto simp: min_def intro!: antisym) lemma arg_min_list_in: "xs \ [] \ arg_min_list f xs \ set xs" by(induction xs rule: induct_list012) (auto simp: Let_def) subsubsection \(In)finiteness\ lemma finite_maxlen: "finite (M::'a list set) \ \n. \s\M. size s < n" proof (induct rule: finite.induct) case emptyI show ?case by simp next case (insertI M xs) then obtain n where "\s\M. length s < n" by blast hence "\s\insert xs M. size s < max n (size xs) + 1" by auto thus ?case .. qed lemma lists_length_Suc_eq: "{xs. set xs \ A \ length xs = Suc n} = (\(xs, n). n#xs) ` ({xs. set xs \ A \ length xs = n} \ A)" by (auto simp: length_Suc_conv) lemma assumes "finite A" shows finite_lists_length_eq: "finite {xs. set xs \ A \ length xs = n}" and card_lists_length_eq: "card {xs. set xs \ A \ length xs = n} = (card A)^n" using \finite A\ by (induct n) (auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong) lemma finite_lists_length_le: assumes "finite A" shows "finite {xs. set xs \ A \ length xs \ n}" (is "finite ?S") proof- have "?S = (\n\{0..n}. {xs. set xs \ A \ length xs = n})" by auto thus ?thesis by (auto intro!: finite_lists_length_eq[OF \finite A\] simp only:) qed lemma card_lists_length_le: assumes "finite A" shows "card {xs. set xs \ A \ length xs \ n} = (\i\n. card A^i)" proof - have "(\i\n. card A^i) = card (\i\n. {xs. set xs \ A \ length xs = i})" using \finite A\ by (subst card_UN_disjoint) (auto simp add: card_lists_length_eq finite_lists_length_eq) also have "(\i\n. {xs. set xs \ A \ length xs = i}) = {xs. set xs \ A \ length xs \ n}" by auto finally show ?thesis by simp qed lemma finite_lists_distinct_length_eq [intro]: assumes "finite A" shows "finite {xs. length xs = n \ distinct xs \ set xs \ A}" (is "finite ?S") proof - have "finite {xs. set xs \ A \ length xs = n}" using \finite A\ by (rule finite_lists_length_eq) moreover have "?S \ {xs. set xs \ A \ length xs = n}" by auto ultimately show ?thesis using finite_subset by auto qed lemma card_lists_distinct_length_eq: assumes "finite A" "k \ card A" shows "card {xs. length xs = k \ distinct xs \ set xs \ A} = \{card A - k + 1 .. card A}" using assms proof (induct k) case 0 then have "{xs. length xs = 0 \ distinct xs \ set xs \ A} = {[]}" by auto then show ?case by simp next case (Suc k) let "?k_list" = "\k xs. length xs = k \ distinct xs \ set xs \ A" have inj_Cons: "\A. inj_on (\(xs, n). n # xs) A" by (rule inj_onI) auto from Suc have "k \ card A" by simp moreover note \finite A\ moreover have "finite {xs. ?k_list k xs}" by (rule finite_subset) (use finite_lists_length_eq[OF \finite A\, of k] in auto) moreover have "\i j. i \ j \ {i} \ (A - set i) \ {j} \ (A - set j) = {}" by auto moreover have "\i. i \ {xs. ?k_list k xs} \ card (A - set i) = card A - k" by (simp add: card_Diff_subset distinct_card) moreover have "{xs. ?k_list (Suc k) xs} = (\(xs, n). n#xs) ` \((\xs. {xs} \ (A - set xs)) ` {xs. ?k_list k xs})" by (auto simp: length_Suc_conv) moreover have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp then have "(card A - k) * \{Suc (card A - k)..card A} = \{Suc (card A - Suc k)..card A}" by (subst prod.insert[symmetric]) (simp add: atLeastAtMost_insertL)+ ultimately show ?case by (simp add: card_image inj_Cons card_UN_disjoint Suc.hyps algebra_simps) qed lemma card_lists_distinct_length_eq': assumes "k < card A" shows "card {xs. length xs = k \ distinct xs \ set xs \ A} = \{card A - k + 1 .. card A}" proof - from \k < card A\ have "finite A" and "k \ card A" using card_infinite by force+ from this show ?thesis by (rule card_lists_distinct_length_eq) qed lemma infinite_UNIV_listI: "\ finite(UNIV::'a list set)" by (metis UNIV_I finite_maxlen length_replicate less_irrefl) subsection \Sorting\ subsubsection \\<^const>\sorted_wrt\\ text \Sometimes the second equation in the definition of \<^const>\sorted_wrt\ is too aggressive because it relates each list element to \emph{all} its successors. Then this equation should be removed and \sorted_wrt2_simps\ should be added instead.\ lemma sorted_wrt1: "sorted_wrt P [x] = True" by(simp) lemma sorted_wrt2: "transp P \ sorted_wrt P (x # y # zs) = (P x y \ sorted_wrt P (y # zs))" proof (induction zs arbitrary: x y) case (Cons z zs) then show ?case by simp (meson transpD)+ qed auto lemmas sorted_wrt2_simps = sorted_wrt1 sorted_wrt2 lemma sorted_wrt_true [simp]: "sorted_wrt (\_ _. True) xs" by (induction xs) simp_all lemma sorted_wrt_append: "sorted_wrt P (xs @ ys) \ sorted_wrt P xs \ sorted_wrt P ys \ (\x\set xs. \y\set ys. P x y)" by (induction xs) auto lemma sorted_wrt_map: "sorted_wrt R (map f xs) = sorted_wrt (\x y. R (f x) (f y)) xs" by (induction xs) simp_all lemma sorted_wrt_rev: "sorted_wrt P (rev xs) = sorted_wrt (\x y. P y x) xs" by (induction xs) (auto simp add: sorted_wrt_append) lemma sorted_wrt_mono_rel: "(\x y. \ x \ set xs; y \ set xs; P x y \ \ Q x y) \ sorted_wrt P xs \ sorted_wrt Q xs" by(induction xs)(auto) lemma sorted_wrt01: "length xs \ 1 \ sorted_wrt P xs" by(auto simp: le_Suc_eq length_Suc_conv) lemma sorted_wrt_iff_nth_less: "sorted_wrt P xs = (\i j. i < j \ j < length xs \ P (xs ! i) (xs ! j))" by (induction xs) (auto simp add: in_set_conv_nth Ball_def nth_Cons split: nat.split) lemma sorted_wrt_nth_less: "\ sorted_wrt P xs; i < j; j < length xs \ \ P (xs ! i) (xs ! j)" by(auto simp: sorted_wrt_iff_nth_less) lemma sorted_wrt_upt[simp]: "sorted_wrt (<) [m..Each element is greater or equal to its index:\ lemma sorted_wrt_less_idx: "sorted_wrt (<) ns \ i < length ns \ i \ ns!i" proof (induction ns arbitrary: i rule: rev_induct) case Nil thus ?case by simp next case snoc thus ?case by (auto simp: nth_append sorted_wrt_append) (metis less_antisym not_less nth_mem) qed subsubsection \\<^const>\sorted\\ context linorder begin text \Sometimes the second equation in the definition of \<^const>\sorted\ is too aggressive because it relates each list element to \emph{all} its successors. Then this equation should be removed and \sorted2_simps\ should be added instead. Executable code is one such use case.\ lemma sorted1: "sorted [x] = True" by simp lemma sorted2: "sorted (x # y # zs) = (x \ y \ sorted (y # zs))" by(induction zs) auto lemmas sorted2_simps = sorted1 sorted2 lemmas [code] = sorted.simps(1) sorted2_simps lemma sorted_append: "sorted (xs@ys) = (sorted xs \ sorted ys \ (\x \ set xs. \y \ set ys. x\y))" by (simp add: sorted_sorted_wrt sorted_wrt_append) lemma sorted_map: "sorted (map f xs) = sorted_wrt (\x y. f x \ f y) xs" by (simp add: sorted_sorted_wrt sorted_wrt_map) lemma sorted01: "length xs \ 1 \ sorted xs" by (simp add: sorted_sorted_wrt sorted_wrt01) lemma sorted_tl: "sorted xs \ sorted (tl xs)" by (cases xs) (simp_all) lemma sorted_iff_nth_mono_less: "sorted xs = (\i j. i < j \ j < length xs \ xs ! i \ xs ! j)" by (simp add: sorted_sorted_wrt sorted_wrt_iff_nth_less) lemma sorted_iff_nth_mono: "sorted xs = (\i j. i \ j \ j < length xs \ xs ! i \ xs ! j)" by (auto simp: sorted_iff_nth_mono_less nat_less_le) lemma sorted_nth_mono: "sorted xs \ i \ j \ j < length xs \ xs!i \ xs!j" by (auto simp: sorted_iff_nth_mono) lemma sorted_rev_nth_mono: "sorted (rev xs) \ i \ j \ j < length xs \ xs!j \ xs!i" using sorted_nth_mono[ of "rev xs" "length xs - j - 1" "length xs - i - 1"] rev_nth[of "length xs - i - 1" "xs"] rev_nth[of "length xs - j - 1" "xs"] by auto lemma sorted_map_remove1: "sorted (map f xs) \ sorted (map f (remove1 x xs))" by (induct xs) (auto) lemma sorted_remove1: "sorted xs \ sorted (remove1 a xs)" using sorted_map_remove1 [of "\x. x"] by simp lemma sorted_butlast: assumes "xs \ []" and "sorted xs" shows "sorted (butlast xs)" proof - from \xs \ []\ obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto with \sorted xs\ show ?thesis by (simp add: sorted_append) qed lemma sorted_replicate [simp]: "sorted(replicate n x)" by(induction n) (auto) lemma sorted_remdups[simp]: "sorted xs \ sorted (remdups xs)" by (induct xs) (auto) lemma sorted_remdups_adj[simp]: "sorted xs \ sorted (remdups_adj xs)" by (induct xs rule: remdups_adj.induct, simp_all split: if_split_asm) lemma sorted_nths: "sorted xs \ sorted (nths xs I)" by(induction xs arbitrary: I)(auto simp: nths_Cons) lemma sorted_distinct_set_unique: assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" "set xs = set ys" shows "xs = ys" proof - from assms have 1: "length xs = length ys" by (auto dest!: distinct_card) from assms show ?thesis proof(induct rule:list_induct2[OF 1]) case 1 show ?case by simp next case 2 thus ?case by simp (metis Diff_insert_absorb antisym insertE insert_iff) qed qed lemma map_sorted_distinct_set_unique: assumes "inj_on f (set xs \ set ys)" assumes "sorted (map f xs)" "distinct (map f xs)" "sorted (map f ys)" "distinct (map f ys)" assumes "set xs = set ys" shows "xs = ys" proof - from assms have "map f xs = map f ys" by (simp add: sorted_distinct_set_unique) with \inj_on f (set xs \ set ys)\ show "xs = ys" by (blast intro: map_inj_on) qed lemma assumes "sorted xs" shows sorted_take: "sorted (take n xs)" and sorted_drop: "sorted (drop n xs)" proof - from assms have "sorted (take n xs @ drop n xs)" by simp then show "sorted (take n xs)" and "sorted (drop n xs)" unfolding sorted_append by simp_all qed lemma sorted_dropWhile: "sorted xs \ sorted (dropWhile P xs)" by (auto dest: sorted_drop simp add: dropWhile_eq_drop) lemma sorted_takeWhile: "sorted xs \ sorted (takeWhile P xs)" by (subst takeWhile_eq_take) (auto dest: sorted_take) lemma sorted_filter: "sorted (map f xs) \ sorted (map f (filter P xs))" by (induct xs) simp_all lemma foldr_max_sorted: assumes "sorted (rev xs)" shows "foldr max xs y = (if xs = [] then y else max (xs ! 0) y)" using assms proof (induct xs) case (Cons x xs) then have "sorted (rev xs)" using sorted_append by auto with Cons show ?case by (cases xs) (auto simp add: sorted_append max_def) qed simp lemma filter_equals_takeWhile_sorted_rev: assumes sorted: "sorted (rev (map f xs))" shows "filter (\x. t < f x) xs = takeWhile (\ x. t < f x) xs" (is "filter ?P xs = ?tW") proof (rule takeWhile_eq_filter[symmetric]) let "?dW" = "dropWhile ?P xs" fix x assume "x \ set ?dW" then obtain i where i: "i < length ?dW" and nth_i: "x = ?dW ! i" unfolding in_set_conv_nth by auto hence "length ?tW + i < length (?tW @ ?dW)" unfolding length_append by simp hence i': "length (map f ?tW) + i < length (map f xs)" by simp have "(map f ?tW @ map f ?dW) ! (length (map f ?tW) + i) \ (map f ?tW @ map f ?dW) ! (length (map f ?tW) + 0)" using sorted_rev_nth_mono[OF sorted _ i', of "length ?tW"] unfolding map_append[symmetric] by simp hence "f x \ f (?dW ! 0)" unfolding nth_append_length_plus nth_i using i preorder_class.le_less_trans[OF le0 i] by simp also have "... \ t" using hd_dropWhile[of "?P" xs] le0[THEN preorder_class.le_less_trans, OF i] using hd_conv_nth[of "?dW"] by simp finally show "\ t < f x" by simp qed lemma sorted_map_same: "sorted (map f (filter (\x. f x = g xs) xs))" proof (induct xs arbitrary: g) case Nil then show ?case by simp next case (Cons x xs) then have "sorted (map f (filter (\y. f y = (\xs. f x) xs) xs))" . moreover from Cons have "sorted (map f (filter (\y. f y = (g \ Cons x) xs) xs))" . ultimately show ?case by simp_all qed lemma sorted_same: "sorted (filter (\x. x = g xs) xs)" using sorted_map_same [of "\x. x"] by simp end lemma sorted_upt[simp]: "sorted [m..Sorting functions\ text\Currently it is not shown that \<^const>\sort\ returns a permutation of its input because the nicest proof is via multisets, which are not part of Main. Alternatively one could define a function that counts the number of occurrences of an element in a list and use that instead of multisets to state the correctness property.\ context linorder begin lemma set_insort_key: "set (insort_key f x xs) = insert x (set xs)" by (induct xs) auto lemma length_insort [simp]: "length (insort_key f x xs) = Suc (length xs)" by (induct xs) simp_all lemma insort_key_left_comm: assumes "f x \ f y" shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)" by (induct xs) (auto simp add: assms dest: antisym) lemma insort_left_comm: "insort x (insort y xs) = insort y (insort x xs)" by (cases "x = y") (auto intro: insort_key_left_comm) lemma comp_fun_commute_insort: "comp_fun_commute insort" proof qed (simp add: insort_left_comm fun_eq_iff) lemma sort_key_simps [simp]: "sort_key f [] = []" "sort_key f (x#xs) = insort_key f x (sort_key f xs)" by (simp_all add: sort_key_def) lemma sort_key_conv_fold: assumes "inj_on f (set xs)" shows "sort_key f xs = fold (insort_key f) xs []" proof - have "fold (insort_key f) (rev xs) = fold (insort_key f) xs" proof (rule fold_rev, rule ext) fix zs fix x y assume "x \ set xs" "y \ set xs" with assms have *: "f y = f x \ y = x" by (auto dest: inj_onD) have **: "x = y \ y = x" by auto show "(insort_key f y \ insort_key f x) zs = (insort_key f x \ insort_key f y) zs" by (induct zs) (auto intro: * simp add: **) qed then show ?thesis by (simp add: sort_key_def foldr_conv_fold) qed lemma sort_conv_fold: "sort xs = fold insort xs []" by (rule sort_key_conv_fold) simp lemma length_sort[simp]: "length (sort_key f xs) = length xs" by (induct xs, auto) lemma set_sort[simp]: "set(sort_key f xs) = set xs" by (induct xs) (simp_all add: set_insort_key) lemma distinct_insort: "distinct (insort_key f x xs) = (x \ set xs \ distinct xs)" by(induct xs)(auto simp: set_insort_key) lemma distinct_sort[simp]: "distinct (sort_key f xs) = distinct xs" by (induct xs) (simp_all add: distinct_insort) lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)" by (induct xs) (auto simp: set_insort_key) lemma sorted_insort: "sorted (insort x xs) = sorted xs" using sorted_insort_key [where f="\x. x"] by simp theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))" by (induct xs) (auto simp:sorted_insort_key) theorem sorted_sort [simp]: "sorted (sort xs)" using sorted_sort_key [where f="\x. x"] by simp lemma insort_not_Nil [simp]: "insort_key f a xs \ []" by (induction xs) simp_all lemma insort_is_Cons: "\x\set xs. f a \ f x \ insort_key f a xs = a # xs" by (cases xs) auto lemma sorted_sort_id: "sorted xs \ sort xs = xs" by (induct xs) (auto simp add: insort_is_Cons) lemma insort_key_remove1: assumes "a \ set xs" and "sorted (map f xs)" and "hd (filter (\x. f a = f x) xs) = a" shows "insort_key f a (remove1 a xs) = xs" using assms proof (induct xs) case (Cons x xs) then show ?case proof (cases "x = a") case False then have "f x \ f a" using Cons.prems by auto then have "f x < f a" using Cons.prems by auto with \f x \ f a\ show ?thesis using Cons by (auto simp: insort_is_Cons) qed (auto simp: insort_is_Cons) qed simp lemma insort_remove1: assumes "a \ set xs" and "sorted xs" shows "insort a (remove1 a xs) = xs" proof (rule insort_key_remove1) define n where "n = length (filter ((=) a) xs) - 1" from \a \ set xs\ show "a \ set xs" . from \sorted xs\ show "sorted (map (\x. x) xs)" by simp from \a \ set xs\ have "a \ set (filter ((=) a) xs)" by auto then have "set (filter ((=) a) xs) \ {}" by auto then have "filter ((=) a) xs \ []" by (auto simp only: set_empty) then have "length (filter ((=) a) xs) > 0" by simp then have n: "Suc n = length (filter ((=) a) xs)" by (simp add: n_def) moreover have "replicate (Suc n) a = a # replicate n a" by simp ultimately show "hd (filter ((=) a) xs) = a" by (simp add: replicate_length_filter) qed lemma finite_sorted_distinct_unique: assumes "finite A" shows "\!xs. set xs = A \ sorted xs \ distinct xs" proof - obtain xs where "distinct xs" "A = set xs" using finite_distinct_list [OF assms] by metis then show ?thesis by (rule_tac a="sort xs" in ex1I) (auto simp: sorted_distinct_set_unique) qed lemma insort_insert_key_triv: "f x \ f ` set xs \ insort_insert_key f x xs = xs" by (simp add: insort_insert_key_def) lemma insort_insert_triv: "x \ set xs \ insort_insert x xs = xs" using insort_insert_key_triv [of "\x. x"] by simp lemma insort_insert_insort_key: "f x \ f ` set xs \ insort_insert_key f x xs = insort_key f x xs" by (simp add: insort_insert_key_def) lemma insort_insert_insort: "x \ set xs \ insort_insert x xs = insort x xs" using insort_insert_insort_key [of "\x. x"] by simp lemma set_insort_insert: "set (insort_insert x xs) = insert x (set xs)" by (auto simp add: insort_insert_key_def set_insort_key) lemma distinct_insort_insert: assumes "distinct xs" shows "distinct (insort_insert_key f x xs)" using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort_key) lemma sorted_insort_insert_key: assumes "sorted (map f xs)" shows "sorted (map f (insort_insert_key f x xs))" using assms by (simp add: insort_insert_key_def sorted_insort_key) lemma sorted_insort_insert: assumes "sorted xs" shows "sorted (insort_insert x xs)" using assms sorted_insort_insert_key [of "\x. x"] by simp lemma filter_insort_triv: "\ P x \ filter P (insort_key f x xs) = filter P xs" by (induct xs) simp_all lemma filter_insort: "sorted (map f xs) \ P x \ filter P (insort_key f x xs) = insort_key f x (filter P xs)" by (induct xs) (auto, subst insort_is_Cons, auto) lemma filter_sort: "filter P (sort_key f xs) = sort_key f (filter P xs)" by (induct xs) (simp_all add: filter_insort_triv filter_insort) lemma remove1_insort [simp]: "remove1 x (insort x xs) = xs" by (induct xs) simp_all end lemma sort_upt [simp]: "sort [m.. \x \ set xs. P x \ List.find P xs = Some (Min {x\set xs. P x})" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases "P x") case True with Cons show ?thesis by (auto intro: Min_eqI [symmetric]) next case False then have "{y. (y = x \ y \ set xs) \ P y} = {y \ set xs. P y}" by auto with Cons False show ?thesis by (simp_all) qed qed lemma sorted_enumerate [simp]: "sorted (map fst (enumerate n xs))" by (simp add: enumerate_eq_zip) text \Stability of \<^const>\sort_key\:\ lemma sort_key_stable: "filter (\y. f y = k) (sort_key f xs) = filter (\y. f y = k) xs" by (induction xs) (auto simp: filter_insort insort_is_Cons filter_insort_triv) corollary stable_sort_key_sort_key: "stable_sort_key sort_key" by(simp add: stable_sort_key_def sort_key_stable) lemma sort_key_const: "sort_key (\x. c) xs = xs" by (metis (mono_tags) filter_True sort_key_stable) subsubsection \\<^const>\transpose\ on sorted lists\ lemma sorted_transpose[simp]: "sorted (rev (map length (transpose xs)))" by (auto simp: sorted_iff_nth_mono rev_nth nth_transpose length_filter_conv_card intro: card_mono) lemma transpose_max_length: "foldr (\xs. max (length xs)) (transpose xs) 0 = length (filter (\x. x \ []) xs)" (is "?L = ?R") proof (cases "transpose xs = []") case False have "?L = foldr max (map length (transpose xs)) 0" by (simp add: foldr_map comp_def) also have "... = length (transpose xs ! 0)" using False sorted_transpose by (simp add: foldr_max_sorted) finally show ?thesis using False by (simp add: nth_transpose) next case True hence "filter (\x. x \ []) xs = []" by (auto intro!: filter_False simp: transpose_empty) thus ?thesis by (simp add: transpose_empty True) qed lemma length_transpose_sorted: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" shows "length (transpose xs) = (if xs = [] then 0 else length (xs ! 0))" proof (cases "xs = []") case False thus ?thesis using foldr_max_sorted[OF sorted] False unfolding length_transpose foldr_map comp_def by simp qed simp lemma nth_nth_transpose_sorted[simp]: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" and i: "i < length (transpose xs)" and j: "j < length (filter (\ys. i < length ys) xs)" shows "transpose xs ! i ! j = xs ! j ! i" using j filter_equals_takeWhile_sorted_rev[OF sorted, of i] nth_transpose[OF i] nth_map[OF j] by (simp add: takeWhile_nth) lemma transpose_column_length: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" and "i < length xs" shows "length (filter (\ys. i < length ys) (transpose xs)) = length (xs ! i)" proof - have "xs \ []" using \i < length xs\ by auto note filter_equals_takeWhile_sorted_rev[OF sorted, simp] { fix j assume "j \ i" note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this \i < length xs\] } note sortedE = this[consumes 1] have "{j. j < length (transpose xs) \ i < length (transpose xs ! j)} = {..< length (xs ! i)}" proof safe fix j assume "j < length (transpose xs)" and "i < length (transpose xs ! j)" with this(2) nth_transpose[OF this(1)] have "i < length (takeWhile (\ys. j < length ys) xs)" by simp from nth_mem[OF this] takeWhile_nth[OF this] show "j < length (xs ! i)" by (auto dest: set_takeWhileD) next fix j assume "j < length (xs ! i)" thus "j < length (transpose xs)" using foldr_max_sorted[OF sorted] \xs \ []\ sortedE[OF le0] by (auto simp: length_transpose comp_def foldr_map) have "Suc i \ length (takeWhile (\ys. j < length ys) xs)" using \i < length xs\ \j < length (xs ! i)\ less_Suc_eq_le by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE) with nth_transpose[OF \j < length (transpose xs)\] show "i < length (transpose xs ! j)" by simp qed thus ?thesis by (simp add: length_filter_conv_card) qed lemma transpose_column: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" and "i < length xs" shows "map (\ys. ys ! i) (filter (\ys. i < length ys) (transpose xs)) = xs ! i" (is "?R = _") proof (rule nth_equalityI) show length: "length ?R = length (xs ! i)" using transpose_column_length[OF assms] by simp fix j assume j: "j < length ?R" note * = less_le_trans[OF this, unfolded length_map, OF length_filter_le] from j have j_less: "j < length (xs ! i)" using length by simp have i_less_tW: "Suc i \ length (takeWhile (\ys. Suc j \ length ys) xs)" proof (rule length_takeWhile_less_P_nth) show "Suc i \ length xs" using \i < length xs\ by simp fix k assume "k < Suc i" hence "k \ i" by auto with sorted_rev_nth_mono[OF sorted this] \i < length xs\ have "length (xs ! i) \ length (xs ! k)" by simp thus "Suc j \ length (xs ! k)" using j_less by simp qed have i_less_filter: "i < length (filter (\ys. j < length ys) xs) " unfolding filter_equals_takeWhile_sorted_rev[OF sorted, of j] using i_less_tW by (simp_all add: Suc_le_eq) from j show "?R ! j = xs ! i ! j" unfolding filter_equals_takeWhile_sorted_rev[OF sorted_transpose, of i] by (simp add: takeWhile_nth nth_nth_transpose_sorted[OF sorted * i_less_filter]) qed lemma transpose_transpose: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" shows "transpose (transpose xs) = takeWhile (\x. x \ []) xs" (is "?L = ?R") proof - have len: "length ?L = length ?R" unfolding length_transpose transpose_max_length using filter_equals_takeWhile_sorted_rev[OF sorted, of 0] by simp { fix i assume "i < length ?R" with less_le_trans[OF _ length_takeWhile_le[of _ xs]] have "i < length xs" by simp } note * = this show ?thesis by (rule nth_equalityI) (simp_all add: len nth_transpose transpose_column[OF sorted] * takeWhile_nth) qed theorem transpose_rectangle: assumes "xs = [] \ n = 0" assumes rect: "\ i. i < length xs \ length (xs ! i) = n" shows "transpose xs = map (\ i. map (\ j. xs ! j ! i) [0..ys. i < length ys) xs = xs" using rect by (auto simp: in_set_conv_nth intro!: filter_True) } ultimately show "\i. i < length (transpose xs) \ ?trans ! i = ?map ! i" by (auto simp: nth_transpose intro: nth_equalityI) qed subsubsection \\sorted_list_of_set\\ text\This function maps (finite) linearly ordered sets to sorted lists. Warning: in most cases it is not a good idea to convert from sets to lists but one should convert in the other direction (via \<^const>\set\).\ context linorder begin definition sorted_list_of_set :: "'a set \ 'a list" where "sorted_list_of_set = folding.F insort []" sublocale sorted_list_of_set: folding insort Nil rewrites "folding.F insort [] = sorted_list_of_set" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show "folding insort" by standard (fact comp_fun_commute) show "folding.F insort [] = sorted_list_of_set" by (simp only: sorted_list_of_set_def) qed lemma sorted_list_of_set_empty: "sorted_list_of_set {} = []" by (fact sorted_list_of_set.empty) lemma sorted_list_of_set_insert [simp]: "finite A \ sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))" by (fact sorted_list_of_set.insert_remove) lemma sorted_list_of_set_eq_Nil_iff [simp]: "finite A \ sorted_list_of_set A = [] \ A = {}" by (auto simp: sorted_list_of_set.remove) lemma set_sorted_list_of_set [simp]: "finite A \ set (sorted_list_of_set A) = A" by(induct A rule: finite_induct) (simp_all add: set_insort_key) lemma sorted_sorted_list_of_set [simp]: "sorted (sorted_list_of_set A)" proof (cases "finite A") case True thus ?thesis by(induction A) (simp_all add: sorted_insort) next case False thus ?thesis by simp qed lemma distinct_sorted_list_of_set [simp]: "distinct (sorted_list_of_set A)" proof (cases "finite A") case True thus ?thesis by(induction A) (simp_all add: distinct_insort) next case False thus ?thesis by simp qed lemmas sorted_list_of_set = set_sorted_list_of_set sorted_sorted_list_of_set distinct_sorted_list_of_set lemma sorted_list_of_set_sort_remdups [code]: "sorted_list_of_set (set xs) = sort (remdups xs)" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_set.eq_fold sort_conv_fold fold_set_fold_remdups) qed lemma sorted_list_of_set_remove: assumes "finite A" shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)" proof (cases "x \ A") case False with assms have "x \ set (sorted_list_of_set A)" by simp with False show ?thesis by (simp add: remove1_idem) next case True then obtain B where A: "A = insert x B" by (rule Set.set_insert) with assms show ?thesis by simp qed end lemma sorted_list_of_set_range [simp]: "sorted_list_of_set {m..\lists\: the list-forming operator over sets\ inductive_set lists :: "'a set => 'a list set" for A :: "'a set" where Nil [intro!, simp]: "[] \ lists A" | Cons [intro!, simp]: "\a \ A; l \ lists A\ \ a#l \ lists A" inductive_cases listsE [elim!]: "x#l \ lists A" inductive_cases listspE [elim!]: "listsp A (x # l)" inductive_simps listsp_simps[code]: "listsp A []" "listsp A (x # xs)" lemma listsp_mono [mono]: "A \ B ==> listsp A \ listsp B" by (rule predicate1I, erule listsp.induct, blast+) lemmas lists_mono = listsp_mono [to_set] lemma listsp_infI: assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l by induct blast+ lemmas lists_IntI = listsp_infI [to_set] lemma listsp_inf_eq [simp]: "listsp (inf A B) = inf (listsp A) (listsp B)" proof (rule mono_inf [where f=listsp, THEN order_antisym]) show "mono listsp" by (simp add: mono_def listsp_mono) show "inf (listsp A) (listsp B) \ listsp (inf A B)" by (blast intro!: listsp_infI) qed lemmas listsp_conj_eq [simp] = listsp_inf_eq [simplified inf_fun_def inf_bool_def] lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set] lemma Cons_in_lists_iff[simp]: "x#xs \ lists A \ x \ A \ xs \ lists A" by auto lemma append_in_listsp_conv [iff]: "(listsp A (xs @ ys)) = (listsp A xs \ listsp A ys)" by (induct xs) auto lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set] lemma in_listsp_conv_set: "(listsp A xs) = (\x \ set xs. A x)" \ \eliminate \listsp\ in favour of \set\\ by (induct xs) auto lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set] lemma in_listspD [dest!]: "listsp A xs ==> \x\set xs. A x" by (rule in_listsp_conv_set [THEN iffD1]) lemmas in_listsD [dest!] = in_listspD [to_set] lemma in_listspI [intro!]: "\x\set xs. A x ==> listsp A xs" by (rule in_listsp_conv_set [THEN iffD2]) lemmas in_listsI [intro!] = in_listspI [to_set] lemma lists_eq_set: "lists A = {xs. set xs \ A}" by auto lemma lists_empty [simp]: "lists {} = {[]}" by auto lemma lists_UNIV [simp]: "lists UNIV = UNIV" by auto lemma lists_image: "lists (f`A) = map f ` lists A" proof - { fix xs have "\x\set xs. x \ f ` A \ xs \ map f ` lists A" by (induct xs) (auto simp del: list.map simp add: list.map[symmetric] intro!: imageI) } then show ?thesis by auto qed subsubsection \Inductive definition for membership\ inductive ListMem :: "'a \ 'a list \ bool" where elem: "ListMem x (x # xs)" | insert: "ListMem x xs \ ListMem x (y # xs)" lemma ListMem_iff: "(ListMem x xs) = (x \ set xs)" proof show "ListMem x xs \ x \ set xs" by (induct set: ListMem) auto show "x \ set xs \ ListMem x xs" by (induct xs) (auto intro: ListMem.intros) qed subsubsection \Lists as Cartesian products\ text\\set_Cons A Xs\: the set of lists with head drawn from \<^term>\A\ and tail drawn from \<^term>\Xs\.\ definition set_Cons :: "'a set \ 'a list set \ 'a list set" where "set_Cons A XS = {z. \x xs. z = x # xs \ x \ A \ xs \ XS}" lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A" by (auto simp add: set_Cons_def) text\Yields the set of lists, all of the same length as the argument and with elements drawn from the corresponding element of the argument.\ primrec listset :: "'a set list \ 'a list set" where "listset [] = {[]}" | "listset (A # As) = set_Cons A (listset As)" subsection \Relations on Lists\ subsubsection \Length Lexicographic Ordering\ text\These orderings preserve well-foundedness: shorter lists precede longer lists. These ordering are not used in dictionaries.\ primrec \ \The lexicographic ordering for lists of the specified length\ lexn :: "('a \ 'a) set \ nat \ ('a list \ 'a list) set" where "lexn r 0 = {}" | "lexn r (Suc n) = (map_prod (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int {(xs, ys). length xs = Suc n \ length ys = Suc n}" definition lex :: "('a \ 'a) set \ ('a list \ 'a list) set" where "lex r = (\n. lexn r n)" \ \Holds only between lists of the same length\ definition lenlex :: "('a \ 'a) set => ('a list \ 'a list) set" where "lenlex r = inv_image (less_than <*lex*> lex r) (\xs. (length xs, xs))" \ \Compares lists by their length and then lexicographically\ lemma wf_lexn: assumes "wf r" shows "wf (lexn r n)" proof (induct n) case (Suc n) have inj: "inj (\(x, xs). x # xs)" using assms by (auto simp: inj_on_def) have wf: "wf (map_prod (\(x, xs). x # xs) (\(x, xs). x # xs) ` (r <*lex*> lexn r n))" by (simp add: Suc.hyps assms wf_lex_prod wf_map_prod_image [OF _ inj]) then show ?case by (rule wf_subset) auto qed auto lemma lexn_length: "(xs, ys) \ lexn r n \ length xs = n \ length ys = n" by (induct n arbitrary: xs ys) auto lemma wf_lex [intro!]: "wf r ==> wf (lex r)" unfolding lex_def apply (rule wf_UN) apply (simp add: wf_lexn) apply (metis DomainE Int_emptyI RangeE lexn_length) done lemma lexn_conv: "lexn r n = {(xs,ys). length xs = n \ length ys = n \ (\xys x y xs' ys'. xs= xys @ x#xs' \ ys= xys @ y # ys' \ (x, y) \ r)}" apply (induct n, simp) apply (simp add: image_Collect lex_prod_def, safe, blast) apply (rule_tac x = "ab # xys" in exI, simp) apply (case_tac xys, simp_all, blast) done text\By Mathias Fleury:\ lemma lexn_transI: assumes "trans r" shows "trans (lexn r n)" unfolding trans_def proof (intro allI impI) fix as bs cs assume asbs: "(as, bs) \ lexn r n" and bscs: "(bs, cs) \ lexn r n" obtain abs a b as' bs' where n: "length as = n" and "length bs = n" and as: "as = abs @ a # as'" and bs: "bs = abs @ b # bs'" and abr: "(a, b) \ r" using asbs unfolding lexn_conv by blast obtain bcs b' c' cs' bs' where n': "length cs = n" and "length bs = n" and bs': "bs = bcs @ b' # bs'" and cs: "cs = bcs @ c' # cs'" and b'c'r: "(b', c') \ r" using bscs unfolding lexn_conv by blast consider (le) "length bcs < length abs" | (eq) "length bcs = length abs" | (ge) "length bcs > length abs" by linarith thus "(as, cs) \ lexn r n" proof cases let ?k = "length bcs" case le hence "as ! ?k = bs ! ?k" unfolding as bs by (simp add: nth_append) hence "(as ! ?k, cs ! ?k) \ r" using b'c'r unfolding bs' cs by auto moreover have "length bcs < length as" using le unfolding as by simp from id_take_nth_drop[OF this] have "as = take ?k as @ as ! ?k # drop (Suc ?k) as" . moreover have "length bcs < length cs" unfolding cs by simp from id_take_nth_drop[OF this] have "cs = take ?k cs @ cs ! ?k # drop (Suc ?k) cs" . moreover have "take ?k as = take ?k cs" using le arg_cong[OF bs, of "take (length bcs)"] unfolding cs as bs' by auto ultimately show ?thesis using n n' unfolding lexn_conv by auto next let ?k = "length abs" case ge hence "bs ! ?k = cs ! ?k" unfolding bs' cs by (simp add: nth_append) hence "(as ! ?k, cs ! ?k) \ r" using abr unfolding as bs by auto moreover have "length abs < length as" using ge unfolding as by simp from id_take_nth_drop[OF this] have "as = take ?k as @ as ! ?k # drop (Suc ?k) as" . moreover have "length abs < length cs" using n n' unfolding as by simp from id_take_nth_drop[OF this] have "cs = take ?k cs @ cs ! ?k # drop (Suc ?k) cs" . moreover have "take ?k as = take ?k cs" using ge arg_cong[OF bs', of "take (length abs)"] unfolding cs as bs by auto ultimately show ?thesis using n n' unfolding lexn_conv by auto next let ?k = "length abs" case eq hence *: "abs = bcs" "b = b'" using bs bs' by auto hence "(a, c') \ r" using abr b'c'r assms unfolding trans_def by blast with * show ?thesis using n n' unfolding lexn_conv as bs cs by auto qed qed lemma lex_conv: "lex r = {(xs,ys). length xs = length ys \ (\xys x y xs' ys'. xs = xys @ x # xs' \ ys = xys @ y # ys' \ (x, y) \ r)}" by (force simp add: lex_def lexn_conv) lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)" by (unfold lenlex_def) blast lemma lenlex_conv: "lenlex r = {(xs,ys). length xs < length ys \ length xs = length ys \ (xs, ys) \ lex r}" by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def) lemma Nil_notin_lex [iff]: "([], ys) \ lex r" by (simp add: lex_conv) lemma Nil2_notin_lex [iff]: "(xs, []) \ lex r" by (simp add:lex_conv) lemma Cons_in_lex [simp]: "((x # xs, y # ys) \ lex r) = ((x, y) \ r \ length xs = length ys \ x = y \ (xs, ys) \ lex r)" apply (simp add: lex_conv) apply (rule iffI) prefer 2 apply (blast intro: Cons_eq_appendI, clarify) by (metis hd_append append_Nil list.sel(1) list.sel(3) tl_append2) lemma lex_append_rightI: "(xs, ys) \ lex r \ length vs = length us \ (xs @ us, ys @ vs) \ lex r" by (fastforce simp: lex_def lexn_conv) lemma lex_append_leftI: "(ys, zs) \ lex r \ (xs @ ys, xs @ zs) \ lex r" by (induct xs) auto lemma lex_append_leftD: "\x. (x,x) \ r \ (xs @ ys, xs @ zs) \ lex r \ (ys, zs) \ lex r" by (induct xs) auto lemma lex_append_left_iff: "\x. (x,x) \ r \ (xs @ ys, xs @ zs) \ lex r \ (ys, zs) \ lex r" by(metis lex_append_leftD lex_append_leftI) lemma lex_take_index: assumes "(xs, ys) \ lex r" obtains i where "i < length xs" and "i < length ys" and "take i xs = take i ys" and "(xs ! i, ys ! i) \ r" proof - obtain n us x xs' y ys' where "(xs, ys) \ lexn r n" and "length xs = n" and "length ys = n" and "xs = us @ x # xs'" and "ys = us @ y # ys'" and "(x, y) \ r" using assms by (fastforce simp: lex_def lexn_conv) then show ?thesis by (intro that [of "length us"]) auto qed subsubsection \Lexicographic Ordering\ text \Classical lexicographic ordering on lists, ie. "a" < "ab" < "b". This ordering does \emph{not} preserve well-foundedness. Author: N. Voelker, March 2005.\ definition lexord :: "('a \ 'a) set \ ('a list \ 'a list) set" where "lexord r = {(x,y). \ a v. y = x @ a # v \ (\ u a b v w. (a,b) \ r \ x = u @ (a # v) \ y = u @ (b # w))}" lemma lexord_Nil_left[simp]: "([],y) \ lexord r = (\ a x. y = a # x)" by (unfold lexord_def, induct_tac y, auto) lemma lexord_Nil_right[simp]: "(x,[]) \ lexord r" by (unfold lexord_def, induct_tac x, auto) lemma lexord_cons_cons[simp]: "((a # x, b # y) \ lexord r) = ((a,b)\ r \ (a = b \ (x,y)\ lexord r))" unfolding lexord_def apply (safe, simp_all) apply (metis hd_append list.sel(1)) apply (metis hd_append list.sel(1) list.sel(3) tl_append2) apply blast by (meson Cons_eq_appendI) lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons lemma lexord_append_rightI: "\ b z. y = b # z \ (x, x @ y) \ lexord r" by (induct_tac x, auto) lemma lexord_append_left_rightI: "(a,b) \ r \ (u @ a # x, u @ b # y) \ lexord r" by (induct_tac u, auto) lemma lexord_append_leftI: " (u,v) \ lexord r \ (x @ u, x @ v) \ lexord r" by (induct x, auto) lemma lexord_append_leftD: "\ (x @ u, x @ v) \ lexord r; (\a. (a,a) \ r) \ \ (u,v) \ lexord r" by (erule rev_mp, induct_tac x, auto) lemma lexord_take_index_conv: "((x,y) \ lexord r) = ((length x < length y \ take (length x) y = x) \ (\i. i < min(length x)(length y) \ take i x = take i y \ (x!i,y!i) \ r))" apply (unfold lexord_def Let_def, clarsimp) apply (rule_tac f = "(% a b. a \ b)" in arg_cong2) apply (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le) apply auto apply (rule_tac x ="length u" in exI, simp) by (metis id_take_nth_drop) \ \lexord is extension of partial ordering List.lex\ lemma lexord_lex: "(x,y) \ lex r = ((x,y) \ lexord r \ length x = length y)" proof (induction x arbitrary: y) case (Cons a x y) then show ?case by (cases y) (force+) qed auto lemma lexord_irreflexive: "\x. (x,x) \ r \ (xs,xs) \ lexord r" by (induct xs) auto text\By Ren\'e Thiemann:\ lemma lexord_partial_trans: "(\x y z. x \ set xs \ (x,y) \ r \ (y,z) \ r \ (x,z) \ r) \ (xs,ys) \ lexord r \ (ys,zs) \ lexord r \ (xs,zs) \ lexord r" proof (induct xs arbitrary: ys zs) case Nil from Nil(3) show ?case unfolding lexord_def by (cases zs, auto) next case (Cons x xs yys zzs) from Cons(3) obtain y ys where yys: "yys = y # ys" unfolding lexord_def by (cases yys, auto) note Cons = Cons[unfolded yys] from Cons(3) have one: "(x,y) \ r \ x = y \ (xs,ys) \ lexord r" by auto from Cons(4) obtain z zs where zzs: "zzs = z # zs" unfolding lexord_def by (cases zzs, auto) note Cons = Cons[unfolded zzs] from Cons(4) have two: "(y,z) \ r \ y = z \ (ys,zs) \ lexord r" by auto { assume "(xs,ys) \ lexord r" and "(ys,zs) \ lexord r" from Cons(1)[OF _ this] Cons(2) have "(xs,zs) \ lexord r" by auto } note ind1 = this { assume "(x,y) \ r" and "(y,z) \ r" from Cons(2)[OF _ this] have "(x,z) \ r" by auto } note ind2 = this from one two ind1 ind2 have "(x,z) \ r \ x = z \ (xs,zs) \ lexord r" by blast thus ?case unfolding zzs by auto qed lemma lexord_trans: "\ (x, y) \ lexord r; (y, z) \ lexord r; trans r \ \ (x, z) \ lexord r" by(auto simp: trans_def intro:lexord_partial_trans) lemma lexord_transI: "trans r \ trans (lexord r)" by (rule transI, drule lexord_trans, blast) lemma lexord_linear: "(\a b. (a,b) \ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" proof (induction x arbitrary: y) case Nil then show ?case by (metis lexord_Nil_left list.exhaust) next case (Cons a x y) then show ?case by (cases y) (force+) -qed +qed lemma lexord_irrefl: "irrefl R \ irrefl (lexord R)" by (simp add: irrefl_def lexord_irreflexive) lemma lexord_asym: assumes "asym R" shows "asym (lexord R)" proof from assms obtain "irrefl R" by (blast elim: asym.cases) then show "irrefl (lexord R)" by (rule lexord_irrefl) next fix xs ys assume "(xs, ys) \ lexord R" then show "(ys, xs) \ lexord R" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then obtain z zs where ys: "ys = z # zs" by (cases ys) auto with assms Cons show ?case by (auto elim: asym.cases) qed qed lemma lexord_asymmetric: assumes "asym R" assumes hyp: "(a, b) \ lexord R" shows "(b, a) \ lexord R" proof - from \asym R\ have "asym (lexord R)" by (rule lexord_asym) then show ?thesis by (rule asym.cases) (auto simp add: hyp) qed text \ Predicate version of lexicographic order integrated with Isabelle's order type classes. Author: Andreas Lochbihler \ context ord begin context notes [[inductive_internals]] begin inductive lexordp :: "'a list \ 'a list \ bool" where Nil: "lexordp [] (y # ys)" | Cons: "x < y \ lexordp (x # xs) (y # ys)" | Cons_eq: "\ \ x < y; \ y < x; lexordp xs ys \ \ lexordp (x # xs) (y # ys)" end lemma lexordp_simps [simp]: "lexordp [] ys = (ys \ [])" "lexordp xs [] = False" "lexordp (x # xs) (y # ys) \ x < y \ \ y < x \ lexordp xs ys" by(subst lexordp.simps, fastforce simp add: neq_Nil_conv)+ inductive lexordp_eq :: "'a list \ 'a list \ bool" where Nil: "lexordp_eq [] ys" | Cons: "x < y \ lexordp_eq (x # xs) (y # ys)" | Cons_eq: "\ \ x < y; \ y < x; lexordp_eq xs ys \ \ lexordp_eq (x # xs) (y # ys)" lemma lexordp_eq_simps [simp]: "lexordp_eq [] ys = True" "lexordp_eq xs [] \ xs = []" "lexordp_eq (x # xs) [] = False" "lexordp_eq (x # xs) (y # ys) \ x < y \ \ y < x \ lexordp_eq xs ys" by(subst lexordp_eq.simps, fastforce)+ lemma lexordp_append_rightI: "ys \ Nil \ lexordp xs (xs @ ys)" by(induct xs)(auto simp add: neq_Nil_conv) lemma lexordp_append_left_rightI: "x < y \ lexordp (us @ x # xs) (us @ y # ys)" by(induct us) auto lemma lexordp_eq_refl: "lexordp_eq xs xs" by(induct xs) simp_all lemma lexordp_append_leftI: "lexordp us vs \ lexordp (xs @ us) (xs @ vs)" by(induct xs) auto lemma lexordp_append_leftD: "\ lexordp (xs @ us) (xs @ vs); \a. \ a < a \ \ lexordp us vs" by(induct xs) auto lemma lexordp_irreflexive: assumes irrefl: "\x. \ x < x" shows "\ lexordp xs xs" proof assume "lexordp xs xs" thus False by(induct xs ys\xs)(simp_all add: irrefl) qed lemma lexordp_into_lexordp_eq: assumes "lexordp xs ys" shows "lexordp_eq xs ys" using assms by induct simp_all end declare ord.lexordp_simps [simp, code] declare ord.lexordp_eq_simps [code, simp] lemma lexord_code [code, code_unfold]: "lexordp = ord.lexordp less" unfolding lexordp_def ord.lexordp_def .. context order begin lemma lexordp_antisym: assumes "lexordp xs ys" "lexordp ys xs" shows False using assms by induct auto lemma lexordp_irreflexive': "\ lexordp xs xs" by(rule lexordp_irreflexive) simp end context linorder begin lemma lexordp_cases [consumes 1, case_names Nil Cons Cons_eq, cases pred: lexordp]: assumes "lexordp xs ys" obtains (Nil) y ys' where "xs = []" "ys = y # ys'" | (Cons) x xs' y ys' where "xs = x # xs'" "ys = y # ys'" "x < y" | (Cons_eq) x xs' ys' where "xs = x # xs'" "ys = x # ys'" "lexordp xs' ys'" using assms by cases (fastforce simp add: not_less_iff_gr_or_eq)+ lemma lexordp_induct [consumes 1, case_names Nil Cons Cons_eq, induct pred: lexordp]: assumes major: "lexordp xs ys" and Nil: "\y ys. P [] (y # ys)" and Cons: "\x xs y ys. x < y \ P (x # xs) (y # ys)" and Cons_eq: "\x xs ys. \ lexordp xs ys; P xs ys \ \ P (x # xs) (x # ys)" shows "P xs ys" using major by induct (simp_all add: Nil Cons not_less_iff_gr_or_eq Cons_eq) lemma lexordp_iff: "lexordp xs ys \ (\x vs. ys = xs @ x # vs) \ (\us a b vs ws. a < b \ xs = us @ a # vs \ ys = us @ b # ws)" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs proof induct case Cons_eq thus ?case by simp (metis append.simps(2)) qed(fastforce intro: disjI2 del: disjCI intro: exI[where x="[]"])+ next assume ?rhs thus ?lhs by(auto intro: lexordp_append_leftI[where us="[]", simplified] lexordp_append_leftI) qed lemma lexordp_conv_lexord: "lexordp xs ys \ (xs, ys) \ lexord {(x, y). x < y}" by(simp add: lexordp_iff lexord_def) lemma lexordp_eq_antisym: assumes "lexordp_eq xs ys" "lexordp_eq ys xs" shows "xs = ys" using assms by induct simp_all lemma lexordp_eq_trans: assumes "lexordp_eq xs ys" and "lexordp_eq ys zs" shows "lexordp_eq xs zs" using assms by (induct arbitrary: zs) (case_tac zs; auto)+ lemma lexordp_trans: assumes "lexordp xs ys" "lexordp ys zs" shows "lexordp xs zs" using assms by (induct arbitrary: zs) (case_tac zs; auto)+ lemma lexordp_linear: "lexordp xs ys \ xs = ys \ lexordp ys xs" by(induct xs arbitrary: ys; case_tac ys; fastforce) lemma lexordp_conv_lexordp_eq: "lexordp xs ys \ lexordp_eq xs ys \ \ lexordp_eq ys xs" (is "?lhs \ ?rhs") proof assume ?lhs hence "\ lexordp_eq ys xs" by induct simp_all with \?lhs\ show ?rhs by (simp add: lexordp_into_lexordp_eq) next assume ?rhs hence "lexordp_eq xs ys" "\ lexordp_eq ys xs" by simp_all thus ?lhs by induct simp_all qed lemma lexordp_eq_conv_lexord: "lexordp_eq xs ys \ xs = ys \ lexordp xs ys" by(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl dest: lexordp_eq_antisym) lemma lexordp_eq_linear: "lexordp_eq xs ys \ lexordp_eq ys xs" by (induct xs arbitrary: ys) (case_tac ys; auto)+ lemma lexordp_linorder: "class.linorder lexordp_eq lexordp" by unfold_locales (auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl lexordp_eq_antisym intro: lexordp_eq_trans del: disjCI intro: lexordp_eq_linear) end lemma sorted_insort_is_snoc: "sorted xs \ \x \ set xs. a \ x \ insort a xs = xs @ [a]" by (induct xs) (auto dest!: insort_is_Cons) subsubsection \Lexicographic combination of measure functions\ text \These are useful for termination proofs\ definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)" lemma wf_measures[simp]: "wf (measures fs)" unfolding measures_def by blast lemma in_measures[simp]: "(x, y) \ measures [] = False" "(x, y) \ measures (f # fs) = (f x < f y \ (f x = f y \ (x, y) \ measures fs))" unfolding measures_def by auto lemma measures_less: "f x < f y ==> (x, y) \ measures (f#fs)" by simp lemma measures_lesseq: "f x \ f y ==> (x, y) \ measures fs ==> (x, y) \ measures (f#fs)" by auto subsubsection \Lifting Relations to Lists: one element\ definition listrel1 :: "('a \ 'a) set \ ('a list \ 'a list) set" where "listrel1 r = {(xs,ys). \us z z' vs. xs = us @ z # vs \ (z,z') \ r \ ys = us @ z' # vs}" lemma listrel1I: "\ (x, y) \ r; xs = us @ x # vs; ys = us @ y # vs \ \ (xs, ys) \ listrel1 r" unfolding listrel1_def by auto lemma listrel1E: "\ (xs, ys) \ listrel1 r; !!x y us vs. \ (x, y) \ r; xs = us @ x # vs; ys = us @ y # vs \ \ P \ \ P" unfolding listrel1_def by auto lemma not_Nil_listrel1 [iff]: "([], xs) \ listrel1 r" unfolding listrel1_def by blast lemma not_listrel1_Nil [iff]: "(xs, []) \ listrel1 r" unfolding listrel1_def by blast lemma Cons_listrel1_Cons [iff]: "(x # xs, y # ys) \ listrel1 r \ (x,y) \ r \ xs = ys \ x = y \ (xs, ys) \ listrel1 r" by (simp add: listrel1_def Cons_eq_append_conv) (blast) lemma listrel1I1: "(x,y) \ r \ (x # xs, y # xs) \ listrel1 r" by fast lemma listrel1I2: "(xs, ys) \ listrel1 r \ (x # xs, x # ys) \ listrel1 r" by fast lemma append_listrel1I: "(xs, ys) \ listrel1 r \ us = vs \ xs = ys \ (us, vs) \ listrel1 r \ (xs @ us, ys @ vs) \ listrel1 r" unfolding listrel1_def by auto (blast intro: append_eq_appendI)+ lemma Cons_listrel1E1[elim!]: assumes "(x # xs, ys) \ listrel1 r" and "\y. ys = y # xs \ (x, y) \ r \ R" and "\zs. ys = x # zs \ (xs, zs) \ listrel1 r \ R" shows R using assms by (cases ys) blast+ lemma Cons_listrel1E2[elim!]: assumes "(xs, y # ys) \ listrel1 r" and "\x. xs = x # ys \ (x, y) \ r \ R" and "\zs. xs = y # zs \ (zs, ys) \ listrel1 r \ R" shows R using assms by (cases xs) blast+ lemma snoc_listrel1_snoc_iff: "(xs @ [x], ys @ [y]) \ listrel1 r \ (xs, ys) \ listrel1 r \ x = y \ xs = ys \ (x,y) \ r" (is "?L \ ?R") proof assume ?L thus ?R by (fastforce simp: listrel1_def snoc_eq_iff_butlast butlast_append) next assume ?R then show ?L unfolding listrel1_def by force qed lemma listrel1_eq_len: "(xs,ys) \ listrel1 r \ length xs = length ys" unfolding listrel1_def by auto lemma listrel1_mono: "r \ s \ listrel1 r \ listrel1 s" unfolding listrel1_def by blast lemma listrel1_converse: "listrel1 (r\) = (listrel1 r)\" unfolding listrel1_def by blast lemma in_listrel1_converse: "(x,y) \ listrel1 (r\) \ (x,y) \ (listrel1 r)\" unfolding listrel1_def by blast lemma listrel1_iff_update: "(xs,ys) \ (listrel1 r) \ (\y n. (xs ! n, y) \ r \ n < length xs \ ys = xs[n:=y])" (is "?L \ ?R") proof assume "?L" then obtain x y u v where "xs = u @ x # v" "ys = u @ y # v" "(x,y) \ r" unfolding listrel1_def by auto then have "ys = xs[length u := y]" and "length u < length xs" and "(xs ! length u, y) \ r" by auto then show "?R" by auto next assume "?R" then obtain x y n where "(xs!n, y) \ r" "n < size xs" "ys = xs[n:=y]" "x = xs!n" by auto then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \ r" by (auto intro: upd_conv_take_nth_drop id_take_nth_drop) then show "?L" by (auto simp: listrel1_def) qed text\Accessible part and wellfoundedness:\ lemma Cons_acc_listrel1I [intro!]: "x \ Wellfounded.acc r \ xs \ Wellfounded.acc (listrel1 r) \ (x # xs) \ Wellfounded.acc (listrel1 r)" apply (induct arbitrary: xs set: Wellfounded.acc) apply (erule thin_rl) apply (erule acc_induct) apply (rule accI) apply (blast) done lemma lists_accD: "xs \ lists (Wellfounded.acc r) \ xs \ Wellfounded.acc (listrel1 r)" proof (induct set: lists) case Nil then show ?case by (meson acc.intros not_listrel1_Nil) next case (Cons a l) then show ?case by blast qed lemma lists_accI: "xs \ Wellfounded.acc (listrel1 r) \ xs \ lists (Wellfounded.acc r)" apply (induct set: Wellfounded.acc) apply clarify apply (rule accI) apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def) done lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r" by (auto simp: wf_acc_iff intro: lists_accD lists_accI[THEN Cons_in_lists_iff[THEN iffD1, THEN conjunct1]]) subsubsection \Lifting Relations to Lists: all elements\ inductive_set listrel :: "('a \ 'b) set \ ('a list \ 'b list) set" for r :: "('a \ 'b) set" where Nil: "([],[]) \ listrel r" | Cons: "[| (x,y) \ r; (xs,ys) \ listrel r |] ==> (x#xs, y#ys) \ listrel r" inductive_cases listrel_Nil1 [elim!]: "([],xs) \ listrel r" inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \ listrel r" inductive_cases listrel_Cons1 [elim!]: "(y#ys,xs) \ listrel r" inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \ listrel r" lemma listrel_eq_len: "(xs, ys) \ listrel r \ length xs = length ys" by(induct rule: listrel.induct) auto lemma listrel_iff_zip [code_unfold]: "(xs,ys) \ listrel r \ length xs = length ys \ (\(x,y) \ set(zip xs ys). (x,y) \ r)" (is "?L \ ?R") proof assume ?L thus ?R by induct (auto intro: listrel_eq_len) next assume ?R thus ?L apply (clarify) by (induct rule: list_induct2) (auto intro: listrel.intros) qed lemma listrel_iff_nth: "(xs,ys) \ listrel r \ length xs = length ys \ (\n < length xs. (xs!n, ys!n) \ r)" (is "?L \ ?R") by (auto simp add: all_set_conv_all_nth listrel_iff_zip) lemma listrel_mono: "r \ s \ listrel r \ listrel s" by (meson listrel_iff_nth subrelI subset_eq) lemma listrel_subset: "r \ A \ A \ listrel r \ lists A \ lists A" apply clarify apply (erule listrel.induct, auto) done lemma listrel_refl_on: "refl_on A r \ refl_on (lists A) (listrel r)" apply (simp add: refl_on_def listrel_subset Ball_def) apply (rule allI) apply (induct_tac x) apply (auto intro: listrel.intros) done lemma listrel_sym: "sym r \ sym (listrel r)" by (simp add: listrel_iff_nth sym_def) lemma listrel_trans: "trans r \ trans (listrel r)" apply (simp add: trans_def) apply (intro allI) apply (rule impI) apply (erule listrel.induct) apply (blast intro: listrel.intros)+ done theorem equiv_listrel: "equiv A r \ equiv (lists A) (listrel r)" by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans) lemma listrel_rtrancl_refl[iff]: "(xs,xs) \ listrel(r\<^sup>*)" using listrel_refl_on[of UNIV, OF refl_rtrancl] by(auto simp: refl_on_def) lemma listrel_rtrancl_trans: "\(xs,ys) \ listrel(r\<^sup>*); (ys,zs) \ listrel(r\<^sup>*)\ \ (xs,zs) \ listrel(r\<^sup>*)" by (metis listrel_trans trans_def trans_rtrancl) lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}" by (blast intro: listrel.intros) lemma listrel_Cons: "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})" by (auto simp add: set_Cons_def intro: listrel.intros) text \Relating \<^term>\listrel1\, \<^term>\listrel\ and closures:\ lemma listrel1_rtrancl_subset_rtrancl_listrel1: "listrel1 (r\<^sup>*) \ (listrel1 r)\<^sup>*" proof (rule subrelI) fix xs ys assume 1: "(xs,ys) \ listrel1 (r\<^sup>*)" { fix x y us vs have "(x,y) \ r\<^sup>* \ (us @ x # vs, us @ y # vs) \ (listrel1 r)\<^sup>*" proof(induct rule: rtrancl.induct) case rtrancl_refl show ?case by simp next case rtrancl_into_rtrancl thus ?case by (metis listrel1I rtrancl.rtrancl_into_rtrancl) qed } thus "(xs,ys) \ (listrel1 r)\<^sup>*" using 1 by(blast elim: listrel1E) qed lemma rtrancl_listrel1_eq_len: "(x,y) \ (listrel1 r)\<^sup>* \ length x = length y" by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len) lemma rtrancl_listrel1_ConsI1: "(xs,ys) \ (listrel1 r)\<^sup>* \ (x#xs,x#ys) \ (listrel1 r)\<^sup>*" apply(induct rule: rtrancl.induct) apply simp by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl) lemma rtrancl_listrel1_ConsI2: "(x,y) \ r\<^sup>* \ (xs, ys) \ (listrel1 r)\<^sup>* \ (x # xs, y # ys) \ (listrel1 r)\<^sup>*" by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1 subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1]) lemma listrel1_subset_listrel: "r \ r' \ refl r' \ listrel1 r \ listrel(r')" by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def) lemma listrel_reflcl_if_listrel1: "(xs,ys) \ listrel1 r \ (xs,ys) \ listrel(r\<^sup>*)" by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip) lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r\<^sup>*) = (listrel1 r)\<^sup>*" proof { fix x y assume "(x,y) \ listrel (r\<^sup>*)" then have "(x,y) \ (listrel1 r)\<^sup>*" by induct (auto intro: rtrancl_listrel1_ConsI2) } then show "listrel (r\<^sup>*) \ (listrel1 r)\<^sup>*" by (rule subrelI) next show "listrel (r\<^sup>*) \ (listrel1 r)\<^sup>*" proof(rule subrelI) fix xs ys assume "(xs,ys) \ (listrel1 r)\<^sup>*" then show "(xs,ys) \ listrel (r\<^sup>*)" proof induct case base show ?case by(auto simp add: listrel_iff_zip set_zip) next case (step ys zs) thus ?case by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans) qed qed qed lemma rtrancl_listrel1_if_listrel: "(xs,ys) \ listrel r \ (xs,ys) \ (listrel1 r)\<^sup>*" by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI) lemma listrel_subset_rtrancl_listrel1: "listrel r \ (listrel1 r)\<^sup>*" by(fast intro:rtrancl_listrel1_if_listrel) subsection \Size function\ lemma [measure_function]: "is_measure f \ is_measure (size_list f)" by (rule is_measure_trivial) lemma [measure_function]: "is_measure f \ is_measure (size_option f)" by (rule is_measure_trivial) lemma size_list_estimation[termination_simp]: "x \ set xs \ y < f x \ y < size_list f xs" by (induct xs) auto lemma size_list_estimation'[termination_simp]: "x \ set xs \ y \ f x \ y \ size_list f xs" by (induct xs) auto lemma size_list_map[simp]: "size_list f (map g xs) = size_list (f \ g) xs" by (induct xs) auto lemma size_list_append[simp]: "size_list f (xs @ ys) = size_list f xs + size_list f ys" by (induct xs, auto) lemma size_list_pointwise[termination_simp]: "(\x. x \ set xs \ f x \ g x) \ size_list f xs \ size_list g xs" by (induct xs) force+ subsection \Monad operation\ definition bind :: "'a list \ ('a \ 'b list) \ 'b list" where "bind xs f = concat (map f xs)" hide_const (open) bind lemma bind_simps [simp]: "List.bind [] f = []" "List.bind (x # xs) f = f x @ List.bind xs f" by (simp_all add: bind_def) lemma list_bind_cong [fundef_cong]: assumes "xs = ys" "(\x. x \ set xs \ f x = g x)" shows "List.bind xs f = List.bind ys g" proof - from assms(2) have "List.bind xs f = List.bind xs g" by (induction xs) simp_all with assms(1) show ?thesis by simp qed lemma set_list_bind: "set (List.bind xs f) = (\x\set xs. set (f x))" by (induction xs) simp_all subsection \Code generation\ text\Optional tail recursive version of \<^const>\map\. Can avoid stack overflow in some target languages.\ fun map_tailrec_rev :: "('a \ 'b) \ 'a list \ 'b list \ 'b list" where "map_tailrec_rev f [] bs = bs" | "map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)" lemma map_tailrec_rev: "map_tailrec_rev f as bs = rev(map f as) @ bs" by(induction as arbitrary: bs) simp_all definition map_tailrec :: "('a \ 'b) \ 'a list \ 'b list" where "map_tailrec f as = rev (map_tailrec_rev f as [])" text\Code equation:\ lemma map_eq_map_tailrec: "map = map_tailrec" by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev) subsubsection \Counterparts for set-related operations\ definition member :: "'a list \ 'a \ bool" where [code_abbrev]: "member xs x \ x \ set xs" text \ Use \member\ only for generating executable code. Otherwise use \<^prop>\x \ set xs\ instead --- it is much easier to reason about. \ lemma member_rec [code]: "member (x # xs) y \ x = y \ member xs y" "member [] y \ False" by (auto simp add: member_def) lemma in_set_member (* FIXME delete candidate *): "x \ set xs \ member xs x" by (simp add: member_def) lemmas list_all_iff [code_abbrev] = fun_cong[OF list.pred_set] definition list_ex :: "('a \ bool) \ 'a list \ bool" where list_ex_iff [code_abbrev]: "list_ex P xs \ Bex (set xs) P" definition list_ex1 :: "('a \ bool) \ 'a list \ bool" where list_ex1_iff [code_abbrev]: "list_ex1 P xs \ (\! x. x \ set xs \ P x)" text \ Usually you should prefer \\x\set xs\, \\x\set xs\ and \\!x. x\set xs \ _\ over \<^const>\list_all\, \<^const>\list_ex\ and \<^const>\list_ex1\ in specifications. \ lemma list_all_simps [code]: "list_all P (x # xs) \ P x \ list_all P xs" "list_all P [] \ True" by (simp_all add: list_all_iff) lemma list_ex_simps [simp, code]: "list_ex P (x # xs) \ P x \ list_ex P xs" "list_ex P [] \ False" by (simp_all add: list_ex_iff) lemma list_ex1_simps [simp, code]: "list_ex1 P [] = False" "list_ex1 P (x # xs) = (if P x then list_all (\y. \ P y \ x = y) xs else list_ex1 P xs)" by (auto simp add: list_ex1_iff list_all_iff) lemma Ball_set_list_all: (* FIXME delete candidate *) "Ball (set xs) P \ list_all P xs" by (simp add: list_all_iff) lemma Bex_set_list_ex: (* FIXME delete candidate *) "Bex (set xs) P \ list_ex P xs" by (simp add: list_ex_iff) lemma list_all_append [simp]: "list_all P (xs @ ys) \ list_all P xs \ list_all P ys" by (auto simp add: list_all_iff) lemma list_ex_append [simp]: "list_ex P (xs @ ys) \ list_ex P xs \ list_ex P ys" by (auto simp add: list_ex_iff) lemma list_all_rev [simp]: "list_all P (rev xs) \ list_all P xs" by (simp add: list_all_iff) lemma list_ex_rev [simp]: "list_ex P (rev xs) \ list_ex P xs" by (simp add: list_ex_iff) lemma list_all_length: "list_all P xs \ (\n < length xs. P (xs ! n))" by (auto simp add: list_all_iff set_conv_nth) lemma list_ex_length: "list_ex P xs \ (\n < length xs. P (xs ! n))" by (auto simp add: list_ex_iff set_conv_nth) lemmas list_all_cong [fundef_cong] = list.pred_cong lemma list_ex_cong [fundef_cong]: "xs = ys \ (\x. x \ set ys \ f x = g x) \ list_ex f xs = list_ex g ys" by (simp add: list_ex_iff) definition can_select :: "('a \ bool) \ 'a set \ bool" where [code_abbrev]: "can_select P A = (\!x\A. P x)" lemma can_select_set_list_ex1 [code]: "can_select P (set A) = list_ex1 P A" by (simp add: list_ex1_iff can_select_def) text \Executable checks for relations on sets\ definition listrel1p :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "listrel1p r xs ys = ((xs, ys) \ listrel1 {(x, y). r x y})" lemma [code_unfold]: "(xs, ys) \ listrel1 r = listrel1p (\x y. (x, y) \ r) xs ys" unfolding listrel1p_def by auto lemma [code]: "listrel1p r [] xs = False" "listrel1p r xs [] = False" "listrel1p r (x # xs) (y # ys) \ r x y \ xs = ys \ x = y \ listrel1p r xs ys" by (simp add: listrel1p_def)+ definition lexordp :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "lexordp r xs ys = ((xs, ys) \ lexord {(x, y). r x y})" lemma [code_unfold]: "(xs, ys) \ lexord r = lexordp (\x y. (x, y) \ r) xs ys" unfolding lexordp_def by auto lemma [code]: "lexordp r xs [] = False" "lexordp r [] (y#ys) = True" "lexordp r (x # xs) (y # ys) = (r x y \ (x = y \ lexordp r xs ys))" unfolding lexordp_def by auto text \Bounded quantification and summation over nats.\ lemma atMost_upto [code_unfold]: "{..n} = set [0..m (\m \ {0..m (\m \ {0..m\n::nat. P m) \ (\m \ {0..n}. P m)" by auto lemma ex_nat_less [code_unfold]: "(\m\n::nat. P m) \ (\m \ {0..n}. P m)" by auto text\Bounded \LEAST\ operator:\ definition "Bleast S P = (LEAST x. x \ S \ P x)" definition "abort_Bleast S P = (LEAST x. x \ S \ P x)" declare [[code abort: abort_Bleast]] lemma Bleast_code [code]: "Bleast (set xs) P = (case filter P (sort xs) of x#xs \ x | [] \ abort_Bleast (set xs) P)" proof (cases "filter P (sort xs)") case Nil thus ?thesis by (simp add: Bleast_def abort_Bleast_def) next case (Cons x ys) have "(LEAST x. x \ set xs \ P x) = x" proof (rule Least_equality) show "x \ set xs \ P x" by (metis Cons Cons_eq_filter_iff in_set_conv_decomp set_sort) next fix y assume "y \ set xs \ P y" hence "y \ set (filter P xs)" by auto thus "x \ y" by (metis Cons eq_iff filter_sort set_ConsD set_sort sorted.simps(2) sorted_sort) qed thus ?thesis using Cons by (simp add: Bleast_def) qed declare Bleast_def[symmetric, code_unfold] text \Summation over ints.\ lemma greaterThanLessThan_upto [code_unfold]: "{i<..Optimizing by rewriting\ definition null :: "'a list \ bool" where [code_abbrev]: "null xs \ xs = []" text \ Efficient emptyness check is implemented by \<^const>\null\. \ lemma null_rec [code]: "null (x # xs) \ False" "null [] \ True" by (simp_all add: null_def) lemma eq_Nil_null: (* FIXME delete candidate *) "xs = [] \ null xs" by (simp add: null_def) lemma equal_Nil_null [code_unfold]: "HOL.equal xs [] \ null xs" "HOL.equal [] = null" by (auto simp add: equal null_def) definition maps :: "('a \ 'b list) \ 'a list \ 'b list" where [code_abbrev]: "maps f xs = concat (map f xs)" definition map_filter :: "('a \ 'b option) \ 'a list \ 'b list" where [code_post]: "map_filter f xs = map (the \ f) (filter (\x. f x \ None) xs)" text \ Operations \<^const>\maps\ and \<^const>\map_filter\ avoid intermediate lists on execution -- do not use for proving. \ lemma maps_simps [code]: "maps f (x # xs) = f x @ maps f xs" "maps f [] = []" by (simp_all add: maps_def) lemma map_filter_simps [code]: "map_filter f (x # xs) = (case f x of None \ map_filter f xs | Some y \ y # map_filter f xs)" "map_filter f [] = []" by (simp_all add: map_filter_def split: option.split) lemma concat_map_maps: (* FIXME delete candidate *) "concat (map f xs) = maps f xs" by (simp add: maps_def) lemma map_filter_map_filter [code_unfold]: "map f (filter P xs) = map_filter (\x. if P x then Some (f x) else None) xs" by (simp add: map_filter_def) text \Optimized code for \\i\{a..b::int}\ and \\n:{a.. and similiarly for \\\.\ definition all_interval_nat :: "(nat \ bool) \ nat \ nat \ bool" where "all_interval_nat P i j \ (\n \ {i.. i \ j \ P i \ all_interval_nat P (Suc i) j" proof - have *: "\n. P i \ \n\{Suc i.. i \ n \ n < j \ P n" proof - fix n assume "P i" "\n\{Suc i.. n" "n < j" then show "P n" by (cases "n = i") simp_all qed show ?thesis by (auto simp add: all_interval_nat_def intro: *) qed lemma list_all_iff_all_interval_nat [code_unfold]: "list_all P [i.. all_interval_nat P i j" by (simp add: list_all_iff all_interval_nat_def) lemma list_ex_iff_not_all_inverval_nat [code_unfold]: "list_ex P [i.. \ (all_interval_nat (Not \ P) i j)" by (simp add: list_ex_iff all_interval_nat_def) definition all_interval_int :: "(int \ bool) \ int \ int \ bool" where "all_interval_int P i j \ (\k \ {i..j}. P k)" lemma [code]: "all_interval_int P i j \ i > j \ P i \ all_interval_int P (i + 1) j" proof - have *: "\k. P i \ \k\{i+1..j}. P k \ i \ k \ k \ j \ P k" proof - fix k assume "P i" "\k\{i+1..j}. P k" "i \ k" "k \ j" then show "P k" by (cases "k = i") simp_all qed show ?thesis by (auto simp add: all_interval_int_def intro: *) qed lemma list_all_iff_all_interval_int [code_unfold]: "list_all P [i..j] \ all_interval_int P i j" by (simp add: list_all_iff all_interval_int_def) lemma list_ex_iff_not_all_inverval_int [code_unfold]: "list_ex P [i..j] \ \ (all_interval_int (Not \ P) i j)" by (simp add: list_ex_iff all_interval_int_def) text \optimized code (tail-recursive) for \<^term>\length\\ definition gen_length :: "nat \ 'a list \ nat" where "gen_length n xs = n + length xs" lemma gen_length_code [code]: "gen_length n [] = n" "gen_length n (x # xs) = gen_length (Suc n) xs" by(simp_all add: gen_length_def) declare list.size(3-4)[code del] lemma length_code [code]: "length = gen_length 0" by(simp add: gen_length_def fun_eq_iff) hide_const (open) member null maps map_filter all_interval_nat all_interval_int gen_length subsubsection \Pretty lists\ ML \ (* Code generation for list literals. *) signature LIST_CODE = sig val add_literal_list: string -> theory -> theory end; structure List_Code : LIST_CODE = struct open Basic_Code_Thingol; fun implode_list t = let fun dest_cons (IConst { sym = Code_Symbol.Constant \<^const_name>\Cons\, ... } `$ t1 `$ t2) = SOME (t1, t2) | dest_cons _ = NONE; val (ts, t') = Code_Thingol.unfoldr dest_cons t; in case t' of IConst { sym = Code_Symbol.Constant \<^const_name>\Nil\, ... } => SOME ts | _ => NONE end; fun print_list (target_fxy, target_cons) pr fxy t1 t2 = Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy ( pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1, Code_Printer.str target_cons, pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2 ); fun add_literal_list target = let fun pretty literals pr _ vars fxy [(t1, _), (t2, _)] = case Option.map (cons t1) (implode_list t2) of SOME ts => Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts) | NONE => print_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2; in Code_Target.set_printings (Code_Symbol.Constant (\<^const_name>\Cons\, [(target, SOME (Code_Printer.complex_const_syntax (2, pretty)))])) end end; \ code_printing type_constructor list \ (SML) "_ list" and (OCaml) "_ list" and (Haskell) "![(_)]" and (Scala) "List[(_)]" | constant Nil \ (SML) "[]" and (OCaml) "[]" and (Haskell) "[]" and (Scala) "!Nil" | class_instance list :: equal \ (Haskell) - | constant "HOL.equal :: 'a list \ 'a list \ bool" \ (Haskell) infix 4 "==" setup \fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell", "Scala"]\ code_reserved SML list code_reserved OCaml list subsubsection \Use convenient predefined operations\ code_printing constant "(@)" \ (SML) infixr 7 "@" and (OCaml) infixr 6 "@" and (Haskell) infixr 5 "++" and (Scala) infixl 7 "++" | constant map \ (Haskell) "map" | constant filter \ (Haskell) "filter" | constant concat \ (Haskell) "concat" | constant List.maps \ (Haskell) "concatMap" | constant rev \ (Haskell) "reverse" | constant zip \ (Haskell) "zip" | constant List.null \ (Haskell) "null" | constant takeWhile \ (Haskell) "takeWhile" | constant dropWhile \ (Haskell) "dropWhile" | constant list_all \ (Haskell) "all" | constant list_ex \ (Haskell) "any" subsubsection \Implementation of sets by lists\ lemma is_empty_set [code]: "Set.is_empty (set xs) \ List.null xs" by (simp add: Set.is_empty_def null_def) lemma empty_set [code]: "{} = set []" by simp lemma UNIV_coset [code]: "UNIV = List.coset []" by simp lemma compl_set [code]: "- set xs = List.coset xs" by simp lemma compl_coset [code]: "- List.coset xs = set xs" by simp lemma [code]: "x \ set xs \ List.member xs x" "x \ List.coset xs \ \ List.member xs x" by (simp_all add: member_def) lemma insert_code [code]: "insert x (set xs) = set (List.insert x xs)" "insert x (List.coset xs) = List.coset (removeAll x xs)" by simp_all lemma remove_code [code]: "Set.remove x (set xs) = set (removeAll x xs)" "Set.remove x (List.coset xs) = List.coset (List.insert x xs)" by (simp_all add: remove_def Compl_insert) lemma filter_set [code]: "Set.filter P (set xs) = set (filter P xs)" by auto lemma image_set [code]: "image f (set xs) = set (map f xs)" by simp lemma subset_code [code]: "set xs \ B \ (\x\set xs. x \ B)" "A \ List.coset ys \ (\y\set ys. y \ A)" "List.coset [] \ set [] \ False" by auto text \A frequent case -- avoid intermediate sets\ lemma [code_unfold]: "set xs \ set ys \ list_all (\x. x \ set ys) xs" by (auto simp: list_all_iff) lemma Ball_set [code]: "Ball (set xs) P \ list_all P xs" by (simp add: list_all_iff) lemma Bex_set [code]: "Bex (set xs) P \ list_ex P xs" by (simp add: list_ex_iff) lemma card_set [code]: "card (set xs) = length (remdups xs)" proof - have "card (set (remdups xs)) = length (remdups xs)" by (rule distinct_card) simp then show ?thesis by simp qed lemma the_elem_set [code]: "the_elem (set [x]) = x" by simp lemma Pow_set [code]: "Pow (set []) = {{}}" "Pow (set (x # xs)) = (let A = Pow (set xs) in A \ insert x ` A)" by (simp_all add: Pow_insert Let_def) definition map_project :: "('a \ 'b option) \ 'a set \ 'b set" where "map_project f A = {b. \ a \ A. f a = Some b}" lemma [code]: "map_project f (set xs) = set (List.map_filter f xs)" by (auto simp add: map_project_def map_filter_def image_def) hide_const (open) map_project text \Operations on relations\ lemma product_code [code]: "Product_Type.product (set xs) (set ys) = set [(x, y). x \ xs, y \ ys]" by (auto simp add: Product_Type.product_def) lemma Id_on_set [code]: "Id_on (set xs) = set [(x, x). x \ xs]" by (auto simp add: Id_on_def) lemma [code]: "R `` S = List.map_project (\(x, y). if x \ S then Some y else None) R" unfolding map_project_def by (auto split: prod.split if_split_asm) lemma trancl_set_ntrancl [code]: "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)" by (simp add: finite_trancl_ntranl) lemma set_relcomp [code]: "set xys O set yzs = set ([(fst xy, snd yz). xy \ xys, yz \ yzs, snd xy = fst yz])" by auto (auto simp add: Bex_def image_def) lemma wf_set [code]: "wf (set xs) = acyclic (set xs)" by (simp add: wf_iff_acyclic_if_finite) subsection \Setup for Lifting/Transfer\ subsubsection \Transfer rules for the Transfer package\ context includes lifting_syntax begin lemma tl_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) tl tl" unfolding tl_def[abs_def] by transfer_prover lemma butlast_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) butlast butlast" by (rule rel_funI, erule list_all2_induct, auto) lemma map_rec: "map f xs = rec_list Nil (%x _ y. Cons (f x) y) xs" by (induct xs) auto lemma append_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A ===> list_all2 A) append append" unfolding List.append_def by transfer_prover lemma rev_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) rev rev" unfolding List.rev_def by transfer_prover lemma filter_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) filter filter" unfolding List.filter_def by transfer_prover lemma fold_transfer [transfer_rule]: "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) fold fold" unfolding List.fold_def by transfer_prover lemma foldr_transfer [transfer_rule]: "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) foldr foldr" unfolding List.foldr_def by transfer_prover lemma foldl_transfer [transfer_rule]: "((B ===> A ===> B) ===> B ===> list_all2 A ===> B) foldl foldl" unfolding List.foldl_def by transfer_prover lemma concat_transfer [transfer_rule]: "(list_all2 (list_all2 A) ===> list_all2 A) concat concat" unfolding List.concat_def by transfer_prover lemma drop_transfer [transfer_rule]: "((=) ===> list_all2 A ===> list_all2 A) drop drop" unfolding List.drop_def by transfer_prover lemma take_transfer [transfer_rule]: "((=) ===> list_all2 A ===> list_all2 A) take take" unfolding List.take_def by transfer_prover lemma list_update_transfer [transfer_rule]: "(list_all2 A ===> (=) ===> A ===> list_all2 A) list_update list_update" unfolding list_update_def by transfer_prover lemma takeWhile_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) takeWhile takeWhile" unfolding takeWhile_def by transfer_prover lemma dropWhile_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) dropWhile dropWhile" unfolding dropWhile_def by transfer_prover lemma zip_transfer [transfer_rule]: "(list_all2 A ===> list_all2 B ===> list_all2 (rel_prod A B)) zip zip" unfolding zip_def by transfer_prover lemma product_transfer [transfer_rule]: "(list_all2 A ===> list_all2 B ===> list_all2 (rel_prod A B)) List.product List.product" unfolding List.product_def by transfer_prover lemma product_lists_transfer [transfer_rule]: "(list_all2 (list_all2 A) ===> list_all2 (list_all2 A)) product_lists product_lists" unfolding product_lists_def by transfer_prover lemma insert_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(A ===> list_all2 A ===> list_all2 A) List.insert List.insert" unfolding List.insert_def [abs_def] by transfer_prover lemma find_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> rel_option A) List.find List.find" unfolding List.find_def by transfer_prover lemma those_transfer [transfer_rule]: "(list_all2 (rel_option P) ===> rel_option (list_all2 P)) those those" unfolding List.those_def by transfer_prover lemma remove1_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(A ===> list_all2 A ===> list_all2 A) remove1 remove1" unfolding remove1_def by transfer_prover lemma removeAll_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(A ===> list_all2 A ===> list_all2 A) removeAll removeAll" unfolding removeAll_def by transfer_prover lemma distinct_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(list_all2 A ===> (=)) distinct distinct" unfolding distinct_def by transfer_prover lemma remdups_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(list_all2 A ===> list_all2 A) remdups remdups" unfolding remdups_def by transfer_prover lemma remdups_adj_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(list_all2 A ===> list_all2 A) remdups_adj remdups_adj" proof (rule rel_funI, erule list_all2_induct) qed (auto simp: remdups_adj_Cons assms[unfolded bi_unique_def] split: list.splits) lemma replicate_transfer [transfer_rule]: "((=) ===> A ===> list_all2 A) replicate replicate" unfolding replicate_def by transfer_prover lemma length_transfer [transfer_rule]: "(list_all2 A ===> (=)) length length" unfolding size_list_overloaded_def size_list_def by transfer_prover lemma rotate1_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) rotate1 rotate1" unfolding rotate1_def by transfer_prover lemma rotate_transfer [transfer_rule]: "((=) ===> list_all2 A ===> list_all2 A) rotate rotate" unfolding rotate_def [abs_def] by transfer_prover lemma nths_transfer [transfer_rule]: "(list_all2 A ===> rel_set (=) ===> list_all2 A) nths nths" unfolding nths_def [abs_def] by transfer_prover lemma subseqs_transfer [transfer_rule]: "(list_all2 A ===> list_all2 (list_all2 A)) subseqs subseqs" unfolding subseqs_def [abs_def] by transfer_prover lemma partition_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> rel_prod (list_all2 A) (list_all2 A)) partition partition" unfolding partition_def by transfer_prover lemma lists_transfer [transfer_rule]: "(rel_set A ===> rel_set (list_all2 A)) lists lists" apply (rule rel_funI, rule rel_setI) apply (erule lists.induct, simp) apply (simp only: rel_set_def list_all2_Cons1, metis lists.Cons) apply (erule lists.induct, simp) apply (simp only: rel_set_def list_all2_Cons2, metis lists.Cons) done lemma set_Cons_transfer [transfer_rule]: "(rel_set A ===> rel_set (list_all2 A) ===> rel_set (list_all2 A)) set_Cons set_Cons" unfolding rel_fun_def rel_set_def set_Cons_def by (fastforce simp add: list_all2_Cons1 list_all2_Cons2) lemma listset_transfer [transfer_rule]: "(list_all2 (rel_set A) ===> rel_set (list_all2 A)) listset listset" unfolding listset_def by transfer_prover lemma null_transfer [transfer_rule]: "(list_all2 A ===> (=)) List.null List.null" unfolding rel_fun_def List.null_def by auto lemma list_all_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> (=)) list_all list_all" unfolding list_all_iff [abs_def] by transfer_prover lemma list_ex_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> (=)) list_ex list_ex" unfolding list_ex_iff [abs_def] by transfer_prover lemma splice_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A ===> list_all2 A) splice splice" apply (rule rel_funI, erule list_all2_induct, simp add: rel_fun_def, simp) apply (rule rel_funI) apply (erule_tac xs=x in list_all2_induct, simp, simp add: rel_fun_def) done lemma shuffles_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A ===> rel_set (list_all2 A)) shuffles shuffles" proof (intro rel_funI, goal_cases) case (1 xs xs' ys ys') thus ?case proof (induction xs ys arbitrary: xs' ys' rule: shuffles.induct) case (3 x xs y ys xs' ys') from "3.prems" obtain x' xs'' where xs': "xs' = x' # xs''" by (cases xs') auto from "3.prems" obtain y' ys'' where ys': "ys' = y' # ys''" by (cases ys') auto have [transfer_rule]: "A x x'" "A y y'" "list_all2 A xs xs''" "list_all2 A ys ys''" using "3.prems" by (simp_all add: xs' ys') have [transfer_rule]: "rel_set (list_all2 A) (shuffles xs (y # ys)) (shuffles xs'' ys')" and [transfer_rule]: "rel_set (list_all2 A) (shuffles (x # xs) ys) (shuffles xs' ys'')" using "3.prems" by (auto intro!: "3.IH" simp: xs' ys') have "rel_set (list_all2 A) ((#) x ` shuffles xs (y # ys) \ (#) y ` shuffles (x # xs) ys) ((#) x' ` shuffles xs'' ys' \ (#) y' ` shuffles xs' ys'')" by transfer_prover thus ?case by (simp add: xs' ys') qed (auto simp: rel_set_def) qed lemma rtrancl_parametric [transfer_rule]: assumes [transfer_rule]: "bi_unique A" "bi_total A" shows "(rel_set (rel_prod A A) ===> rel_set (rel_prod A A)) rtrancl rtrancl" unfolding rtrancl_def by transfer_prover lemma monotone_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone" unfolding monotone_def[abs_def] by transfer_prover lemma fun_ord_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total C" shows "((A ===> B ===> (=)) ===> (C ===> A) ===> (C ===> B) ===> (=)) fun_ord fun_ord" unfolding fun_ord_def[abs_def] by transfer_prover lemma fun_lub_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_unique A" shows "((rel_set A ===> B) ===> rel_set (C ===> A) ===> C ===> B) fun_lub fun_lub" unfolding fun_lub_def[abs_def] by transfer_prover 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,852 +1,859 @@ (* 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 + 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 end diff --git a/src/HOL/ROOT b/src/HOL/ROOT --- a/src/HOL/ROOT +++ b/src/HOL/ROOT @@ -1,1127 +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 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,2043 +1,2054 @@ (* 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) -> ({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) * (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 (** 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 wit_goals = maps (BNF_Def.mk_wit_goals var_as var_as' sets) Iwits; val lost_wits = 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})); 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 {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_rel, 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 (_, abs_G, rep_G) = strip3 quot_thm; + 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); val wit_thms = (case wit_thmss of [] => wit_thms_of_bnf bnf_F | _ => wit_thmss); (* 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); (* val wits_G = [abs_G o wit_F] *) val wits_G = map (fn (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) Iwits; (* 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]; 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 = 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); 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 Quotient_Info.lookup_quotients lthy absT_name of - SOME qs => quotient_bnf qs + 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); in which_bnf input_thm 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); 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)) #> 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)) #> 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) (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); 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); 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; diff --git a/src/HOL/Transitive_Closure.thy b/src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy +++ b/src/HOL/Transitive_Closure.thy @@ -1,1291 +1,1345 @@ (* Title: HOL/Transitive_Closure.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section \Reflexive and Transitive closure of a relation\ theory Transitive_Closure imports Relation abbrevs "^*" = "\<^sup>*" "\<^sup>*\<^sup>*" and "^+" = "\<^sup>+" "\<^sup>+\<^sup>+" and "^=" = "\<^sup>=" "\<^sup>=\<^sup>=" begin ML_file \~~/src/Provers/trancl.ML\ text \ \rtrancl\ is reflexive/transitive closure, \trancl\ is transitive closure, \reflcl\ is reflexive closure. These postfix operators have \<^emph>\maximum priority\, forcing their operands to be atomic. \ context notes [[inductive_internals]] begin inductive_set rtrancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>*)" [1000] 999) for r :: "('a \ 'a) set" where rtrancl_refl [intro!, Pure.intro!, simp]: "(a, a) \ r\<^sup>*" | rtrancl_into_rtrancl [Pure.intro]: "(a, b) \ r\<^sup>* \ (b, c) \ r \ (a, c) \ r\<^sup>*" inductive_set trancl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>+)" [1000] 999) for r :: "('a \ 'a) set" where r_into_trancl [intro, Pure.intro]: "(a, b) \ r \ (a, b) \ r\<^sup>+" | trancl_into_trancl [Pure.intro]: "(a, b) \ r\<^sup>+ \ (b, c) \ r \ (a, c) \ r\<^sup>+" notation rtranclp ("(_\<^sup>*\<^sup>*)" [1000] 1000) and tranclp ("(_\<^sup>+\<^sup>+)" [1000] 1000) declare rtrancl_def [nitpick_unfold del] rtranclp_def [nitpick_unfold del] trancl_def [nitpick_unfold del] tranclp_def [nitpick_unfold del] end abbreviation reflcl :: "('a \ 'a) set \ ('a \ 'a) set" ("(_\<^sup>=)" [1000] 999) where "r\<^sup>= \ r \ Id" abbreviation reflclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" ("(_\<^sup>=\<^sup>=)" [1000] 1000) where "r\<^sup>=\<^sup>= \ sup r (=)" notation (ASCII) rtrancl ("(_^*)" [1000] 999) and trancl ("(_^+)" [1000] 999) and reflcl ("(_^=)" [1000] 999) and rtranclp ("(_^**)" [1000] 1000) and tranclp ("(_^++)" [1000] 1000) and reflclp ("(_^==)" [1000] 1000) subsection \Reflexive closure\ lemma refl_reflcl[simp]: "refl (r\<^sup>=)" by (simp add: refl_on_def) lemma antisym_reflcl[simp]: "antisym (r\<^sup>=) = antisym r" by (simp add: antisym_def) lemma trans_reflclI[simp]: "trans r \ trans (r\<^sup>=)" unfolding trans_def by blast lemma reflclp_idemp [simp]: "(P\<^sup>=\<^sup>=)\<^sup>=\<^sup>= = P\<^sup>=\<^sup>=" by blast subsection \Reflexive-transitive closure\ lemma reflcl_set_eq [pred_set_conv]: "(sup (\x y. (x, y) \ r) (=)) = (\x y. (x, y) \ r \ Id)" by (auto simp: fun_eq_iff) lemma r_into_rtrancl [intro]: "\p. p \ r \ p \ r\<^sup>*" \ \\rtrancl\ of \r\ contains \r\\ apply (simp only: split_tupled_all) apply (erule rtrancl_refl [THEN rtrancl_into_rtrancl]) done lemma r_into_rtranclp [intro]: "r x y \ r\<^sup>*\<^sup>* x y" \ \\rtrancl\ of \r\ contains \r\\ by (erule rtranclp.rtrancl_refl [THEN rtranclp.rtrancl_into_rtrancl]) lemma rtranclp_mono: "r \ s \ r\<^sup>*\<^sup>* \ s\<^sup>*\<^sup>*" \ \monotonicity of \rtrancl\\ apply (rule predicate2I) apply (erule rtranclp.induct) apply (rule_tac [2] rtranclp.rtrancl_into_rtrancl, blast+) done lemma mono_rtranclp[mono]: "(\a b. x a b \ y a b) \ x\<^sup>*\<^sup>* a b \ y\<^sup>*\<^sup>* a b" using rtranclp_mono[of x y] by auto lemmas rtrancl_mono = rtranclp_mono [to_set] theorem rtranclp_induct [consumes 1, case_names base step, induct set: rtranclp]: assumes a: "r\<^sup>*\<^sup>* a b" and cases: "P a" "\y z. r\<^sup>*\<^sup>* a y \ r y z \ P y \ P z" shows "P b" using a by (induct x\a b) (rule cases)+ lemmas rtrancl_induct [induct set: rtrancl] = rtranclp_induct [to_set] lemmas rtranclp_induct2 = rtranclp_induct[of _ "(ax,ay)" "(bx,by)", split_rule, consumes 1, case_names refl step] lemmas rtrancl_induct2 = rtrancl_induct[of "(ax,ay)" "(bx,by)", split_format (complete), consumes 1, case_names refl step] lemma refl_rtrancl: "refl (r\<^sup>*)" unfolding refl_on_def by fast text \Transitivity of transitive closure.\ lemma trans_rtrancl: "trans (r\<^sup>*)" proof (rule transI) fix x y z assume "(x, y) \ r\<^sup>*" assume "(y, z) \ r\<^sup>*" then show "(x, z) \ r\<^sup>*" proof induct case base show "(x, y) \ r\<^sup>*" by fact next case (step u v) from \(x, u) \ r\<^sup>*\ and \(u, v) \ r\ show "(x, v) \ r\<^sup>*" .. qed qed lemmas rtrancl_trans = trans_rtrancl [THEN transD] lemma rtranclp_trans: assumes "r\<^sup>*\<^sup>* x y" and "r\<^sup>*\<^sup>* y z" shows "r\<^sup>*\<^sup>* x z" using assms(2,1) by induct iprover+ lemma rtranclE [cases set: rtrancl]: fixes a b :: 'a assumes major: "(a, b) \ r\<^sup>*" obtains (base) "a = b" | (step) y where "(a, y) \ r\<^sup>*" and "(y, b) \ r" \ \elimination of \rtrancl\ -- by induction on a special formula\ proof - have "a = b \ (\y. (a, y) \ r\<^sup>* \ (y, b) \ r)" by (rule major [THEN rtrancl_induct]) blast+ then show ?thesis by (auto intro: base step) qed lemma rtrancl_Int_subset: "Id \ s \ (r\<^sup>* \ s) O r \ s \ r\<^sup>* \ s" apply clarify apply (erule rtrancl_induct, auto) done lemma converse_rtranclp_into_rtranclp: "r a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>*\<^sup>* a c" by (rule rtranclp_trans) iprover+ lemmas converse_rtrancl_into_rtrancl = converse_rtranclp_into_rtranclp [to_set] text \\<^medskip> More \<^term>\r\<^sup>*\ equations and inclusions.\ lemma rtranclp_idemp [simp]: "(r\<^sup>*\<^sup>*)\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*" apply (auto intro!: order_antisym) apply (erule rtranclp_induct) apply (rule rtranclp.rtrancl_refl) apply (blast intro: rtranclp_trans) done lemmas rtrancl_idemp [simp] = rtranclp_idemp [to_set] lemma rtrancl_idemp_self_comp [simp]: "R\<^sup>* O R\<^sup>* = R\<^sup>*" apply (rule set_eqI) apply (simp only: split_tupled_all) apply (blast intro: rtrancl_trans) done lemma rtrancl_subset_rtrancl: "r \ s\<^sup>* \ r\<^sup>* \ s\<^sup>*" by (drule rtrancl_mono, simp) lemma rtranclp_subset: "R \ S \ S \ R\<^sup>*\<^sup>* \ S\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*" apply (drule rtranclp_mono) apply (drule rtranclp_mono, simp) done lemmas rtrancl_subset = rtranclp_subset [to_set] lemma rtranclp_sup_rtranclp: "(sup (R\<^sup>*\<^sup>*) (S\<^sup>*\<^sup>*))\<^sup>*\<^sup>* = (sup R S)\<^sup>*\<^sup>*" by (blast intro!: rtranclp_subset intro: rtranclp_mono [THEN predicate2D]) lemmas rtrancl_Un_rtrancl = rtranclp_sup_rtranclp [to_set] lemma rtranclp_reflclp [simp]: "(R\<^sup>=\<^sup>=)\<^sup>*\<^sup>* = R\<^sup>*\<^sup>*" by (blast intro!: rtranclp_subset) lemmas rtrancl_reflcl [simp] = rtranclp_reflclp [to_set] lemma rtrancl_r_diff_Id: "(r - Id)\<^sup>* = r\<^sup>*" by (rule rtrancl_subset [symmetric]) auto lemma rtranclp_r_diff_Id: "(inf r (\))\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*" by (rule rtranclp_subset [symmetric]) auto theorem rtranclp_converseD: assumes "(r\\)\<^sup>*\<^sup>* x y" shows "r\<^sup>*\<^sup>* y x" using assms by induct (iprover intro: rtranclp_trans dest!: conversepD)+ lemmas rtrancl_converseD = rtranclp_converseD [to_set] theorem rtranclp_converseI: assumes "r\<^sup>*\<^sup>* y x" shows "(r\\)\<^sup>*\<^sup>* x y" using assms by induct (iprover intro: rtranclp_trans conversepI)+ lemmas rtrancl_converseI = rtranclp_converseI [to_set] lemma rtrancl_converse: "(r\)\<^sup>* = (r\<^sup>*)\" by (fast dest!: rtrancl_converseD intro!: rtrancl_converseI) lemma sym_rtrancl: "sym r \ sym (r\<^sup>*)" by (simp only: sym_conv_converse_eq rtrancl_converse [symmetric]) theorem converse_rtranclp_induct [consumes 1, case_names base step]: assumes major: "r\<^sup>*\<^sup>* a b" and cases: "P b" "\y z. r y z \ r\<^sup>*\<^sup>* z b \ P z \ P y" shows "P a" using rtranclp_converseI [OF major] by induct (iprover intro: cases dest!: conversepD rtranclp_converseD)+ lemmas converse_rtrancl_induct = converse_rtranclp_induct [to_set] lemmas converse_rtranclp_induct2 = converse_rtranclp_induct [of _ "(ax, ay)" "(bx, by)", split_rule, consumes 1, case_names refl step] lemmas converse_rtrancl_induct2 = converse_rtrancl_induct [of "(ax, ay)" "(bx, by)", split_format (complete), consumes 1, case_names refl step] lemma converse_rtranclpE [consumes 1, case_names base step]: assumes major: "r\<^sup>*\<^sup>* x z" and cases: "x = z \ P" "\y. r x y \ r\<^sup>*\<^sup>* y z \ P" shows P proof - have "x = z \ (\y. r x y \ r\<^sup>*\<^sup>* y z)" by (rule_tac major [THEN converse_rtranclp_induct]) iprover+ then show ?thesis by (auto intro: cases) qed lemmas converse_rtranclE = converse_rtranclpE [to_set] lemmas converse_rtranclpE2 = converse_rtranclpE [of _ "(xa,xb)" "(za,zb)", split_rule] lemmas converse_rtranclE2 = converse_rtranclE [of "(xa,xb)" "(za,zb)", split_rule] lemma r_comp_rtrancl_eq: "r O r\<^sup>* = r\<^sup>* O r" by (blast elim: rtranclE converse_rtranclE intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl) lemma rtrancl_unfold: "r\<^sup>* = Id \ r\<^sup>* O r" by (auto intro: rtrancl_into_rtrancl elim: rtranclE) lemma rtrancl_Un_separatorE: "(a, b) \ (P \ Q)\<^sup>* \ \x y. (a, x) \ P\<^sup>* \ (x, y) \ Q \ x = y \ (a, b) \ P\<^sup>*" proof (induct rule: rtrancl.induct) case rtrancl_refl then show ?case by blast next case rtrancl_into_rtrancl then show ?case by (blast intro: rtrancl_trans) qed lemma rtrancl_Un_separator_converseE: "(a, b) \ (P \ Q)\<^sup>* \ \x y. (x, b) \ P\<^sup>* \ (y, x) \ Q \ y = x \ (a, b) \ P\<^sup>*" proof (induct rule: converse_rtrancl_induct) case base then show ?case by blast next case step then show ?case by (blast intro: rtrancl_trans) qed lemma Image_closed_trancl: assumes "r `` X \ X" shows "r\<^sup>* `` X = X" proof - from assms have **: "{y. \x\X. (x, y) \ r} \ X" by auto have "x \ X" if 1: "(y, x) \ r\<^sup>*" and 2: "y \ X" for x y proof - from 1 show "x \ X" proof induct case base show ?case by (fact 2) next case step with ** show ?case by auto qed qed then show ?thesis by auto qed subsection \Transitive closure\ lemma trancl_mono: "\p. p \ r\<^sup>+ \ r \ s \ p \ s\<^sup>+" apply (simp add: split_tupled_all) apply (erule trancl.induct) apply (iprover dest: subsetD)+ done lemma r_into_trancl': "\p. p \ r \ p \ r\<^sup>+" by (simp only: split_tupled_all) (erule r_into_trancl) text \\<^medskip> Conversions between \trancl\ and \rtrancl\.\ lemma tranclp_into_rtranclp: "r\<^sup>+\<^sup>+ a b \ r\<^sup>*\<^sup>* a b" by (erule tranclp.induct) iprover+ lemmas trancl_into_rtrancl = tranclp_into_rtranclp [to_set] lemma rtranclp_into_tranclp1: assumes "r\<^sup>*\<^sup>* a b" shows "r b c \ r\<^sup>+\<^sup>+ a c" using assms by (induct arbitrary: c) iprover+ lemmas rtrancl_into_trancl1 = rtranclp_into_tranclp1 [to_set] lemma rtranclp_into_tranclp2: "r a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>+\<^sup>+ a c" \ \intro rule from \r\ and \rtrancl\\ apply (erule rtranclp.cases, iprover) apply (rule rtranclp_trans [THEN rtranclp_into_tranclp1]) apply (simp | rule r_into_rtranclp)+ done lemmas rtrancl_into_trancl2 = rtranclp_into_tranclp2 [to_set] text \Nice induction rule for \trancl\\ lemma tranclp_induct [consumes 1, case_names base step, induct pred: tranclp]: assumes a: "r\<^sup>+\<^sup>+ a b" and cases: "\y. r a y \ P y" "\y z. r\<^sup>+\<^sup>+ a y \ r y z \ P y \ P z" shows "P b" using a by (induct x\a b) (iprover intro: cases)+ lemmas trancl_induct [induct set: trancl] = tranclp_induct [to_set] lemmas tranclp_induct2 = tranclp_induct [of _ "(ax, ay)" "(bx, by)", split_rule, consumes 1, case_names base step] lemmas trancl_induct2 = trancl_induct [of "(ax, ay)" "(bx, by)", split_format (complete), consumes 1, case_names base step] lemma tranclp_trans_induct: assumes major: "r\<^sup>+\<^sup>+ x y" and cases: "\x y. r x y \ P x y" "\x y z. r\<^sup>+\<^sup>+ x y \ P x y \ r\<^sup>+\<^sup>+ y z \ P y z \ P x z" shows "P x y" \ \Another induction rule for trancl, incorporating transitivity\ by (iprover intro: major [THEN tranclp_induct] cases) lemmas trancl_trans_induct = tranclp_trans_induct [to_set] lemma tranclE [cases set: trancl]: assumes "(a, b) \ r\<^sup>+" obtains (base) "(a, b) \ r" | (step) c where "(a, c) \ r\<^sup>+" and "(c, b) \ r" using assms by cases simp_all lemma trancl_Int_subset: "r \ s \ (r\<^sup>+ \ s) O r \ s \ r\<^sup>+ \ s" apply clarify apply (erule trancl_induct, auto) done lemma trancl_unfold: "r\<^sup>+ = r \ r\<^sup>+ O r" by (auto intro: trancl_into_trancl elim: tranclE) text \Transitivity of \<^term>\r\<^sup>+\\ lemma trans_trancl [simp]: "trans (r\<^sup>+)" proof (rule transI) fix x y z assume "(x, y) \ r\<^sup>+" assume "(y, z) \ r\<^sup>+" then show "(x, z) \ r\<^sup>+" proof induct case (base u) from \(x, y) \ r\<^sup>+\ and \(y, u) \ r\ show "(x, u) \ r\<^sup>+" .. next case (step u v) from \(x, u) \ r\<^sup>+\ and \(u, v) \ r\ show "(x, v) \ r\<^sup>+" .. qed qed lemmas trancl_trans = trans_trancl [THEN transD] lemma tranclp_trans: assumes "r\<^sup>+\<^sup>+ x y" and "r\<^sup>+\<^sup>+ y z" shows "r\<^sup>+\<^sup>+ x z" using assms(2,1) by induct iprover+ lemma trancl_id [simp]: "trans r \ r\<^sup>+ = r" apply auto apply (erule trancl_induct, assumption) apply (unfold trans_def, blast) done lemma rtranclp_tranclp_tranclp: assumes "r\<^sup>*\<^sup>* x y" shows "\z. r\<^sup>+\<^sup>+ y z \ r\<^sup>+\<^sup>+ x z" using assms by induct (iprover intro: tranclp_trans)+ lemmas rtrancl_trancl_trancl = rtranclp_tranclp_tranclp [to_set] lemma tranclp_into_tranclp2: "r a b \ r\<^sup>+\<^sup>+ b c \ r\<^sup>+\<^sup>+ a c" by (erule tranclp_trans [OF tranclp.r_into_trancl]) lemmas trancl_into_trancl2 = tranclp_into_tranclp2 [to_set] lemma tranclp_converseI: "(r\<^sup>+\<^sup>+)\\ x y \ (r\\)\<^sup>+\<^sup>+ x y" apply (drule conversepD) apply (erule tranclp_induct) apply (iprover intro: conversepI tranclp_trans)+ done lemmas trancl_converseI = tranclp_converseI [to_set] lemma tranclp_converseD: "(r\\)\<^sup>+\<^sup>+ x y \ (r\<^sup>+\<^sup>+)\\ x y" apply (rule conversepI) apply (erule tranclp_induct) apply (iprover dest: conversepD intro: tranclp_trans)+ done lemmas trancl_converseD = tranclp_converseD [to_set] lemma tranclp_converse: "(r\\)\<^sup>+\<^sup>+ = (r\<^sup>+\<^sup>+)\\" by (fastforce simp add: fun_eq_iff intro!: tranclp_converseI dest!: tranclp_converseD) lemmas trancl_converse = tranclp_converse [to_set] lemma sym_trancl: "sym r \ sym (r\<^sup>+)" by (simp only: sym_conv_converse_eq trancl_converse [symmetric]) lemma converse_tranclp_induct [consumes 1, case_names base step]: assumes major: "r\<^sup>+\<^sup>+ a b" and cases: "\y. r y b \ P y" "\y z. r y z \ r\<^sup>+\<^sup>+ z b \ P z \ P y" shows "P a" apply (rule tranclp_induct [OF tranclp_converseI, OF conversepI, OF major]) apply (blast intro: cases) apply (blast intro: assms dest!: tranclp_converseD) done lemmas converse_trancl_induct = converse_tranclp_induct [to_set] lemma tranclpD: "R\<^sup>+\<^sup>+ x y \ \z. R x z \ R\<^sup>*\<^sup>* z y" apply (erule converse_tranclp_induct, auto) apply (blast intro: rtranclp_trans) done lemmas tranclD = tranclpD [to_set] lemma converse_tranclpE: assumes major: "tranclp r x z" and base: "r x z \ P" and step: "\y. r x y \ tranclp r y z \ P" shows P proof - from tranclpD [OF major] obtain y where "r x y" and "rtranclp r y z" by iprover from this(2) show P proof (cases rule: rtranclp.cases) case rtrancl_refl with \r x y\ base show P by iprover next case rtrancl_into_rtrancl from this have "tranclp r y z" by (iprover intro: rtranclp_into_tranclp1) with \r x y\ step show P by iprover qed qed lemmas converse_tranclE = converse_tranclpE [to_set] lemma tranclD2: "(x, y) \ R\<^sup>+ \ \z. (x, z) \ R\<^sup>* \ (z, y) \ R" by (blast elim: tranclE intro: trancl_into_rtrancl) lemma irrefl_tranclI: "r\ \ r\<^sup>* = {} \ (x, x) \ r\<^sup>+" by (blast elim: tranclE dest: trancl_into_rtrancl) lemma irrefl_trancl_rD: "\x. (x, x) \ r\<^sup>+ \ (x, y) \ r \ x \ y" by (blast dest: r_into_trancl) lemma trancl_subset_Sigma_aux: "(a, b) \ r\<^sup>* \ r \ A \ A \ a = b \ a \ A" by (induct rule: rtrancl_induct) auto lemma trancl_subset_Sigma: "r \ A \ A \ r\<^sup>+ \ A \ A" apply (clarsimp simp:) apply (erule tranclE) apply (blast dest!: trancl_into_rtrancl trancl_subset_Sigma_aux)+ done lemma reflclp_tranclp [simp]: "(r\<^sup>+\<^sup>+)\<^sup>=\<^sup>= = r\<^sup>*\<^sup>*" apply (safe intro!: order_antisym) apply (erule tranclp_into_rtranclp) apply (blast elim: rtranclp.cases dest: rtranclp_into_tranclp1) done lemmas reflcl_trancl [simp] = reflclp_tranclp [to_set] lemma trancl_reflcl [simp]: "(r\<^sup>=)\<^sup>+ = r\<^sup>*" proof - have "(a, b) \ (r\<^sup>=)\<^sup>+ \ (a, b) \ r\<^sup>*" for a b by (force dest: trancl_into_rtrancl) moreover have "(a, b) \ (r\<^sup>=)\<^sup>+" if "(a, b) \ r\<^sup>*" for a b using that proof (cases a b rule: rtranclE) case step show ?thesis by (rule rtrancl_into_trancl1) (use step in auto) qed auto - ultimately show ?thesis + ultimately show ?thesis by auto qed lemma rtrancl_trancl_reflcl [code]: "r\<^sup>* = (r\<^sup>+)\<^sup>=" by simp lemma trancl_empty [simp]: "{}\<^sup>+ = {}" by (auto elim: trancl_induct) lemma rtrancl_empty [simp]: "{}\<^sup>* = Id" by (rule subst [OF reflcl_trancl]) simp lemma rtranclpD: "R\<^sup>*\<^sup>* a b \ a = b \ a \ b \ R\<^sup>+\<^sup>+ a b" by (force simp: reflclp_tranclp [symmetric] simp del: reflclp_tranclp) lemmas rtranclD = rtranclpD [to_set] lemma rtrancl_eq_or_trancl: "(x,y) \ R\<^sup>* \ x = y \ x \ y \ (x, y) \ R\<^sup>+" by (fast elim: trancl_into_rtrancl dest: rtranclD) lemma trancl_unfold_right: "r\<^sup>+ = r\<^sup>* O r" by (auto dest: tranclD2 intro: rtrancl_into_trancl1) lemma trancl_unfold_left: "r\<^sup>+ = r O r\<^sup>*" by (auto dest: tranclD intro: rtrancl_into_trancl2) lemma trancl_insert: "(insert (y, x) r)\<^sup>+ = r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*}" \ \primitive recursion for \trancl\ over finite relations\ proof - have "\a b. (a, b) \ (insert (y, x) r)\<^sup>+ \ (a, b) \ r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*}" by (erule trancl_induct) (blast intro: rtrancl_into_trancl1 trancl_into_rtrancl trancl_trans)+ moreover have "r\<^sup>+ \ {(a, b). (a, y) \ r\<^sup>* \ (x, b) \ r\<^sup>*} \ (insert (y, x) r)\<^sup>+" by (blast intro: trancl_mono rtrancl_mono [THEN [2] rev_subsetD] rtrancl_trancl_trancl rtrancl_into_trancl2) ultimately show ?thesis by auto qed lemma trancl_insert2: "(insert (a, b) r)\<^sup>+ = r\<^sup>+ \ {(x, y). ((x, a) \ r\<^sup>+ \ x = a) \ ((b, y) \ r\<^sup>+ \ y = b)}" by (auto simp: trancl_insert rtrancl_eq_or_trancl) lemma rtrancl_insert: "(insert (a,b) r)\<^sup>* = r\<^sup>* \ {(x, y). (x, a) \ r\<^sup>* \ (b, y) \ r\<^sup>*}" using trancl_insert[of a b r] by (simp add: rtrancl_trancl_reflcl del: reflcl_trancl) blast text \Simplifying nested closures\ lemma rtrancl_trancl_absorb[simp]: "(R\<^sup>*)\<^sup>+ = R\<^sup>*" by (simp add: trans_rtrancl) lemma trancl_rtrancl_absorb[simp]: "(R\<^sup>+)\<^sup>* = R\<^sup>*" by (subst reflcl_trancl[symmetric]) simp lemma rtrancl_reflcl_absorb[simp]: "(R\<^sup>*)\<^sup>= = R\<^sup>*" by auto text \\Domain\ and \Range\\ lemma Domain_rtrancl [simp]: "Domain (R\<^sup>*) = UNIV" by blast lemma Range_rtrancl [simp]: "Range (R\<^sup>*) = UNIV" by blast lemma rtrancl_Un_subset: "(R\<^sup>* \ S\<^sup>*) \ (R \ S)\<^sup>*" by (rule rtrancl_Un_rtrancl [THEN subst]) fast lemma in_rtrancl_UnI: "x \ R\<^sup>* \ x \ S\<^sup>* \ x \ (R \ S)\<^sup>*" by (blast intro: subsetD [OF rtrancl_Un_subset]) lemma trancl_domain [simp]: "Domain (r\<^sup>+) = Domain r" by (unfold Domain_unfold) (blast dest: tranclD) lemma trancl_range [simp]: "Range (r\<^sup>+) = Range r" unfolding Domain_converse [symmetric] by (simp add: trancl_converse [symmetric]) -lemma Not_Domain_rtrancl: +lemma Not_Domain_rtrancl: assumes "x \ Domain R" shows "(x, y) \ R\<^sup>* \ x = y" proof - have "(x, y) \ R\<^sup>* \ x = y" by (erule rtrancl_induct) (use assms in auto) then show ?thesis by auto qed lemma trancl_subset_Field2: "r\<^sup>+ \ Field r \ Field r" apply clarify apply (erule trancl_induct) apply (auto simp: Field_def) done lemma finite_trancl[simp]: "finite (r\<^sup>+) = finite r" proof show "finite (r\<^sup>+) \ finite r" by (blast intro: r_into_trancl' finite_subset) show "finite r \ finite (r\<^sup>+)" apply (rule trancl_subset_Field2 [THEN finite_subset]) apply (auto simp: finite_Field) done qed lemma finite_rtrancl_Image[simp]: assumes "finite R" "finite A" shows "finite (R\<^sup>* `` A)" proof (rule ccontr) assume "infinite (R\<^sup>* `` A)" with assms show False by(simp add: rtrancl_trancl_reflcl Un_Image del: reflcl_trancl) qed text \More about converse \rtrancl\ and \trancl\, should be merged with main body.\ lemma single_valued_confluent: assumes "single_valued r" and xy: "(x, y) \ r\<^sup>*" and xz: "(x, z) \ r\<^sup>*" shows "(y, z) \ r\<^sup>* \ (z, y) \ r\<^sup>*" using xy proof (induction rule: rtrancl_induct) case base show ?case - by (simp add: assms) + by (simp add: assms) next case (step y z) with xz \single_valued r\ show ?case apply (auto simp: elim: converse_rtranclE dest: single_valuedD) apply (blast intro: rtrancl_trans) done qed lemma r_r_into_trancl: "(a, b) \ R \ (b, c) \ R \ (a, c) \ R\<^sup>+" by (fast intro: trancl_trans) lemma trancl_into_trancl: "(a, b) \ r\<^sup>+ \ (b, c) \ r \ (a, c) \ r\<^sup>+" by (induct rule: trancl_induct) (fast intro: r_r_into_trancl trancl_trans)+ lemma tranclp_rtranclp_tranclp: "r\<^sup>+\<^sup>+ a b \ r\<^sup>*\<^sup>* b c \ r\<^sup>+\<^sup>+ a c" apply (drule tranclpD) apply (elim exE conjE) apply (drule rtranclp_trans, assumption) apply (drule (2) rtranclp_into_tranclp2) done +lemma rtranclp_conversep: "r\\\<^sup>*\<^sup>* = r\<^sup>*\<^sup>*\\" + by(auto simp add: fun_eq_iff intro: rtranclp_converseI rtranclp_converseD) + +lemmas symp_rtranclp = sym_rtrancl[to_pred] + +lemmas symp_conv_conversep_eq = sym_conv_converse_eq[to_pred] + +lemmas rtranclp_tranclp_absorb [simp] = rtrancl_trancl_absorb[to_pred] +lemmas tranclp_rtranclp_absorb [simp] = trancl_rtrancl_absorb[to_pred] +lemmas rtranclp_reflclp_absorb [simp] = rtrancl_reflcl_absorb[to_pred] + lemmas trancl_rtrancl_trancl = tranclp_rtranclp_tranclp [to_set] lemmas transitive_closure_trans [trans] = r_r_into_trancl trancl_trans rtrancl_trans trancl.trancl_into_trancl trancl_into_trancl2 rtrancl.rtrancl_into_rtrancl converse_rtrancl_into_rtrancl rtrancl_trancl_trancl trancl_rtrancl_trancl lemmas transitive_closurep_trans' [trans] = tranclp_trans rtranclp_trans tranclp.trancl_into_trancl tranclp_into_tranclp2 rtranclp.rtrancl_into_rtrancl converse_rtranclp_into_rtranclp rtranclp_tranclp_tranclp tranclp_rtranclp_tranclp declare trancl_into_rtrancl [elim] +subsection \Symmetric closure\ + +definition symclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" +where "symclp r x y \ r x y \ r y x" + +lemma symclpI [simp, intro?]: + shows symclpI1: "r x y \ symclp r x y" + and symclpI2: "r y x \ symclp r x y" +by(simp_all add: symclp_def) + +lemma symclpE [consumes 1, cases pred]: + assumes "symclp r x y" + obtains (base) "r x y" | (sym) "r y x" +using assms by(auto simp add: symclp_def) + +lemma symclp_pointfree: "symclp r = sup r r\\" + by(auto simp add: symclp_def fun_eq_iff) + +lemma symclp_greater: "r \ symclp r" + by(simp add: symclp_pointfree) + +lemma symclp_conversep [simp]: "symclp r\\ = symclp r" + by(simp add: symclp_pointfree sup.commute) + +lemma symp_symclp [simp]: "symp (symclp r)" + by(auto simp add: symp_def elim: symclpE intro: symclpI) + +lemma symp_symclp_eq: "symp r \ symclp r = r" + by(simp add: symclp_pointfree symp_conv_conversep_eq) + +lemma symp_rtranclp_symclp [simp]: "symp (symclp r)\<^sup>*\<^sup>*" + by(simp add: symp_rtranclp) + +lemma rtranclp_symclp_sym [sym]: "(symclp r)\<^sup>*\<^sup>* x y \ (symclp r)\<^sup>*\<^sup>* y x" + by(rule sympD[OF symp_rtranclp_symclp]) + +lemma symclp_idem [simp]: "symclp (symclp r) = symclp r" + by(simp add: symclp_pointfree sup_commute converse_join) + +lemma reflp_rtranclp [simp]: "reflp R\<^sup>*\<^sup>*" + using refl_rtrancl[to_pred, of R] reflp_refl_eq[of "{(x, y). R\<^sup>*\<^sup>* x y}"] by simp subsection \The power operation on relations\ text \\R ^^ n = R O \ O R\, the n-fold composition of \R\\ overloading relpow \ "compow :: nat \ ('a \ 'a) set \ ('a \ 'a) set" relpowp \ "compow :: nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" begin primrec relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where "relpow 0 R = Id" | "relpow (Suc n) R = (R ^^ n) O R" primrec relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" where "relpowp 0 R = HOL.eq" | "relpowp (Suc n) R = (R ^^ n) OO R" end lemma relpowp_relpow_eq [pred_set_conv]: "(\x y. (x, y) \ R) ^^ n = (\x y. (x, y) \ R ^^ n)" for R :: "'a rel" by (induct n) (simp_all add: relcompp_relcomp_eq) text \For code generation:\ definition relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where relpow_code_def [code_abbrev]: "relpow = compow" definition relpowp :: "nat \ ('a \ 'a \ bool) \ ('a \ 'a \ bool)" where relpowp_code_def [code_abbrev]: "relpowp = compow" lemma [code]: "relpow (Suc n) R = (relpow n R) O R" "relpow 0 R = Id" by (simp_all add: relpow_code_def) lemma [code]: "relpowp (Suc n) R = (R ^^ n) OO R" "relpowp 0 R = HOL.eq" by (simp_all add: relpowp_code_def) hide_const (open) relpow hide_const (open) relpowp lemma relpow_1 [simp]: "R ^^ 1 = R" for R :: "('a \ 'a) set" by simp lemma relpowp_1 [simp]: "P ^^ 1 = P" for P :: "'a \ 'a \ bool" by (fact relpow_1 [to_pred]) lemma relpow_0_I: "(x, x) \ R ^^ 0" by simp lemma relpowp_0_I: "(P ^^ 0) x x" by (fact relpow_0_I [to_pred]) lemma relpow_Suc_I: "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" by auto lemma relpowp_Suc_I: "(P ^^ n) x y \ P y z \ (P ^^ Suc n) x z" by (fact relpow_Suc_I [to_pred]) lemma relpow_Suc_I2: "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" by (induct n arbitrary: z) (simp, fastforce) lemma relpowp_Suc_I2: "P x y \ (P ^^ n) y z \ (P ^^ Suc n) x z" by (fact relpow_Suc_I2 [to_pred]) lemma relpow_0_E: "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" by simp lemma relpowp_0_E: "(P ^^ 0) x y \ (x = y \ Q) \ Q" by (fact relpow_0_E [to_pred]) lemma relpow_Suc_E: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" by auto lemma relpowp_Suc_E: "(P ^^ Suc n) x z \ (\y. (P ^^ n) x y \ P y z \ Q) \ Q" by (fact relpow_Suc_E [to_pred]) lemma relpow_E: "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) \ (\y m. n = Suc m \ (x, y) \ R ^^ m \ (y, z) \ R \ P) \ P" by (cases n) auto lemma relpowp_E: "(P ^^ n) x z \ (n = 0 \ x = z \ Q) \ (\y m. n = Suc m \ (P ^^ m) x y \ P y z \ Q) \ Q" by (fact relpow_E [to_pred]) lemma relpow_Suc_D2: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" by (induct n arbitrary: x z) (blast intro: relpow_0_I relpow_Suc_I elim: relpow_0_E relpow_Suc_E)+ lemma relpowp_Suc_D2: "(P ^^ Suc n) x z \ \y. P x y \ (P ^^ n) y z" by (fact relpow_Suc_D2 [to_pred]) lemma relpow_Suc_E2: "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n \ P) \ P" by (blast dest: relpow_Suc_D2) lemma relpowp_Suc_E2: "(P ^^ Suc n) x z \ (\y. P x y \ (P ^^ n) y z \ Q) \ Q" by (fact relpow_Suc_E2 [to_pred]) lemma relpow_Suc_D2': "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" by (induct n) (simp_all, blast) lemma relpowp_Suc_D2': "\x y z. (P ^^ n) x y \ P y z \ (\w. P x w \ (P ^^ n) w z)" by (fact relpow_Suc_D2' [to_pred]) lemma relpow_E2: assumes "(x, z) \ R ^^ n" "n = 0 \ x = z \ P" "\y m. n = Suc m \ (x, y) \ R \ (y, z) \ R ^^ m \ P" shows "P" proof (cases n) case 0 with assms show ?thesis by simp next case (Suc m) with assms relpow_Suc_D2' [of m R] show ?thesis by force qed lemma relpowp_E2: "(P ^^ n) x z \ (n = 0 \ x = z \ Q) \ (\y m. n = Suc m \ P x y \ (P ^^ m) y z \ Q) \ Q" by (fact relpow_E2 [to_pred]) lemma relpow_add: "R ^^ (m + n) = R^^m O R^^n" by (induct n) auto lemma relpowp_add: "P ^^ (m + n) = P ^^ m OO P ^^ n" by (fact relpow_add [to_pred]) lemma relpow_commute: "R O R ^^ n = R ^^ n O R" by (induct n) (simp_all add: O_assoc [symmetric]) lemma relpowp_commute: "P OO P ^^ n = P ^^ n OO P" by (fact relpow_commute [to_pred]) lemma relpow_empty: "0 < n \ ({} :: ('a \ 'a) set) ^^ n = {}" by (cases n) auto lemma relpowp_bot: "0 < n \ (\ :: 'a \ 'a \ bool) ^^ n = \" by (fact relpow_empty [to_pred]) lemma rtrancl_imp_UN_relpow: assumes "p \ R\<^sup>*" shows "p \ (\n. R ^^ n)" proof (cases p) case (Pair x y) with assms have "(x, y) \ R\<^sup>*" by simp then have "(x, y) \ (\n. R ^^ n)" proof induct case base show ?case by (blast intro: relpow_0_I) next case step then show ?case by (blast intro: relpow_Suc_I) qed with Pair show ?thesis by simp qed lemma rtranclp_imp_Sup_relpowp: assumes "(P\<^sup>*\<^sup>*) x y" shows "(\n. P ^^ n) x y" using assms and rtrancl_imp_UN_relpow [of "(x, y)", to_pred] by simp lemma relpow_imp_rtrancl: assumes "p \ R ^^ n" shows "p \ R\<^sup>*" proof (cases p) case (Pair x y) with assms have "(x, y) \ R ^^ n" by simp then have "(x, y) \ R\<^sup>*" proof (induct n arbitrary: x y) case 0 then show ?case by simp next case Suc then show ?case by (blast elim: relpow_Suc_E intro: rtrancl_into_rtrancl) qed with Pair show ?thesis by simp qed lemma relpowp_imp_rtranclp: "(P ^^ n) x y \ (P\<^sup>*\<^sup>*) x y" using relpow_imp_rtrancl [of "(x, y)", to_pred] by simp lemma rtrancl_is_UN_relpow: "R\<^sup>* = (\n. R ^^ n)" by (blast intro: rtrancl_imp_UN_relpow relpow_imp_rtrancl) lemma rtranclp_is_Sup_relpowp: "P\<^sup>*\<^sup>* = (\n. P ^^ n)" using rtrancl_is_UN_relpow [to_pred, of P] by auto lemma rtrancl_power: "p \ R\<^sup>* \ (\n. p \ R ^^ n)" by (simp add: rtrancl_is_UN_relpow) lemma rtranclp_power: "(P\<^sup>*\<^sup>*) x y \ (\n. (P ^^ n) x y)" by (simp add: rtranclp_is_Sup_relpowp) lemma trancl_power: "p \ R\<^sup>+ \ (\n > 0. p \ R ^^ n)" proof - have "((a, b) \ R\<^sup>+) = (\n>0. (a, b) \ R ^^ n)" for a b proof safe show "(a, b) \ R\<^sup>+ \ \n>0. (a, b) \ R ^^ n" apply (drule tranclD2) apply (fastforce simp: rtrancl_is_UN_relpow relcomp_unfold) done show "(a, b) \ R\<^sup>+" if "n > 0" "(a, b) \ R ^^ n" for n proof (cases n) case (Suc m) with that show ?thesis by (auto simp: dest: relpow_imp_rtrancl rtrancl_into_trancl1) qed (use that in auto) qed then show ?thesis by (cases p) auto qed lemma tranclp_power: "(P\<^sup>+\<^sup>+) x y \ (\n > 0. (P ^^ n) x y)" using trancl_power [to_pred, of P "(x, y)"] by simp lemma rtrancl_imp_relpow: "p \ R\<^sup>* \ \n. p \ R ^^ n" by (auto dest: rtrancl_imp_UN_relpow) lemma rtranclp_imp_relpowp: "(P\<^sup>*\<^sup>*) x y \ \n. (P ^^ n) x y" by (auto dest: rtranclp_imp_Sup_relpowp) text \By Sternagel/Thiemann:\ lemma relpow_fun_conv: "(a, b) \ R ^^ n \ (\f. f 0 = a \ f n = b \ (\i R))" proof (induct n arbitrary: b) case 0 show ?case by auto next case (Suc n) show ?case proof (simp add: relcomp_unfold Suc) show "(\y. (\f. f 0 = a \ f n = y \ (\i R)) \ (y,b) \ R) \ (\f. f 0 = a \ f(Suc n) = b \ (\i R))" (is "?l = ?r") proof assume ?l then obtain c f where 1: "f 0 = a" "f n = c" "\i. i < n \ (f i, f (Suc i)) \ R" "(c,b) \ R" by auto let ?g = "\ m. if m = Suc n then b else f m" show ?r by (rule exI[of _ ?g]) (simp add: 1) next assume ?r then obtain f where 1: "f 0 = a" "b = f (Suc n)" "\i. i < Suc n \ (f i, f (Suc i)) \ R" by auto show ?l by (rule exI[of _ "f n"], rule conjI, rule exI[of _ f], insert 1, auto) qed qed qed lemma relpowp_fun_conv: "(P ^^ n) x y \ (\f. f 0 = x \ f n = y \ (\i 'a) set" assumes "finite R" and "k > 0" shows "R^^k \ (\n\{n. 0 < n \ n \ card R}. R^^n)" (is "_ \ ?r") proof - have "(a, b) \ R^^(Suc k) \ \n. 0 < n \ n \ card R \ (a, b) \ R^^n" for a b k proof (induct k arbitrary: b) case 0 then have "R \ {}" by auto with card_0_eq[OF \finite R\] have "card R \ Suc 0" by auto then show ?case using 0 by force next case (Suc k) then obtain a' where "(a, a') \ R^^(Suc k)" and "(a', b) \ R" by auto from Suc(1)[OF \(a, a') \ R^^(Suc k)\] obtain n where "n \ card R" and "(a, a') \ R ^^ n" by auto have "(a, b) \ R^^(Suc n)" using \(a, a') \ R^^n\ and \(a', b)\ R\ by auto from \n \ card R\ consider "n < card R" | "n = card R" by force then show ?case proof cases case 1 then show ?thesis using \(a, b) \ R^^(Suc n)\ Suc_leI[OF \n < card R\] by blast next case 2 from \(a, b) \ R ^^ (Suc n)\ [unfolded relpow_fun_conv] obtain f where "f 0 = a" and "f (Suc n) = b" and steps: "\i. i \ n \ (f i, f (Suc i)) \ R" by auto let ?p = "\i. (f i, f(Suc i))" let ?N = "{i. i \ n}" have "?p ` ?N \ R" using steps by auto from card_mono[OF assms(1) this] have "card (?p ` ?N) \ card R" . also have "\ < card ?N" using \n = card R\ by simp finally have "\ inj_on ?p ?N" by (rule pigeonhole) then obtain i j where i: "i \ n" and j: "j \ n" and ij: "i \ j" and pij: "?p i = ?p j" by (auto simp: inj_on_def) let ?i = "min i j" let ?j = "max i j" have i: "?i \ n" and j: "?j \ n" and pij: "?p ?i = ?p ?j" and ij: "?i < ?j" using i j ij pij unfolding min_def max_def by auto from i j pij ij obtain i j where i: "i \ n" and j: "j \ n" and ij: "i < j" and pij: "?p i = ?p j" by blast let ?g = "\l. if l \ i then f l else f (l + (j - i))" let ?n = "Suc (n - (j - i))" have abl: "(a, b) \ R ^^ ?n" unfolding relpow_fun_conv proof (rule exI[of _ ?g], intro conjI impI allI) show "?g ?n = b" using \f(Suc n) = b\ j ij by auto next fix k assume "k < ?n" show "(?g k, ?g (Suc k)) \ R" proof (cases "k < i") case True with i have "k \ n" by auto from steps[OF this] show ?thesis using True by simp next case False then have "i \ k" by auto show ?thesis proof (cases "k = i") case True then show ?thesis using ij pij steps[OF i] by simp next case False with \i \ k\ have "i < k" by auto then have small: "k + (j - i) \ n" using \k by arith show ?thesis using steps[OF small] \i by auto qed qed qed (simp add: \f 0 = a\) moreover have "?n \ n" using i j ij by arith ultimately show ?thesis using \n = card R\ by blast qed qed then show ?thesis using gr0_implies_Suc[OF \k > 0\] by auto qed lemma relpow_finite_bounded: fixes R :: "('a \ 'a) set" assumes "finite R" shows "R^^k \ (\n\{n. n \ card R}. R^^n)" apply (cases k, force) apply (use relpow_finite_bounded1[OF assms, of k] in auto) done lemma rtrancl_finite_eq_relpow: "finite R \ R\<^sup>* = (\n\{n. n \ card R}. R^^n)" by (fastforce simp: rtrancl_power dest: relpow_finite_bounded) lemma trancl_finite_eq_relpow: "finite R \ R\<^sup>+ = (\n\{n. 0 < n \ n \ card R}. R^^n)" apply (auto simp: trancl_power) apply (auto dest: relpow_finite_bounded1) done lemma finite_relcomp[simp,intro]: assumes "finite R" and "finite S" shows "finite (R O S)" proof- have "R O S = (\(x, y)\R. \(u, v)\S. if u = y then {(x, v)} else {})" by (force simp: split_def image_constant_conv split: if_splits) then show ?thesis using assms by clarsimp qed lemma finite_relpow [simp, intro]: fixes R :: "('a \ 'a) set" assumes "finite R" shows "n > 0 \ finite (R^^n)" proof (induct n) case 0 then show ?case by simp next case (Suc n) then show ?case by (cases n) (use assms in simp_all) qed lemma single_valued_relpow: fixes R :: "('a \ 'a) set" shows "single_valued R \ single_valued (R ^^ n)" proof (induct n arbitrary: R) case 0 then show ?case by simp next case (Suc n) show ?case by (rule single_valuedI) (use Suc in \fast dest: single_valuedD elim: relpow_Suc_E\) qed subsection \Bounded transitive closure\ definition ntrancl :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where "ntrancl n R = (\i\{i. 0 < i \ i \ Suc n}. R ^^ i)" lemma ntrancl_Zero [simp, code]: "ntrancl 0 R = R" proof show "R \ ntrancl 0 R" unfolding ntrancl_def by fastforce have "0 < i \ i \ Suc 0 \ i = 1" for i by auto then show "ntrancl 0 R \ R" unfolding ntrancl_def by auto qed lemma ntrancl_Suc [simp]: "ntrancl (Suc n) R = ntrancl n R O (Id \ R)" proof have "(a, b) \ ntrancl n R O (Id \ R)" if "(a, b) \ ntrancl (Suc n) R" for a b proof - from that obtain i where "0 < i" "i \ Suc (Suc n)" "(a, b) \ R ^^ i" unfolding ntrancl_def by auto show ?thesis proof (cases "i = 1") case True from this \(a, b) \ R ^^ i\ show ?thesis by (auto simp: ntrancl_def) next case False with \0 < i\ obtain j where j: "i = Suc j" "0 < j" by (cases i) auto with \(a, b) \ R ^^ i\ obtain c where c1: "(a, c) \ R ^^ j" and c2: "(c, b) \ R" by auto from c1 j \i \ Suc (Suc n)\ have "(a, c) \ ntrancl n R" by (fastforce simp: ntrancl_def) with c2 show ?thesis by fastforce qed qed then show "ntrancl (Suc n) R \ ntrancl n R O (Id \ R)" by auto show "ntrancl n R O (Id \ R) \ ntrancl (Suc n) R" by (fastforce simp: ntrancl_def) qed lemma [code]: "ntrancl (Suc n) r = (let r' = ntrancl n r in r' \ r' O r)" by (auto simp: Let_def) lemma finite_trancl_ntranl: "finite R \ trancl R = ntrancl (card R - 1) R" by (cases "card R") (auto simp: trancl_finite_eq_relpow relpow_empty ntrancl_def) subsection \Acyclic relations\ definition acyclic :: "('a \ 'a) set \ bool" where "acyclic r \ (\x. (x,x) \ r\<^sup>+)" abbreviation acyclicP :: "('a \ 'a \ bool) \ bool" where "acyclicP r \ acyclic {(x, y). r x y}" lemma acyclic_irrefl [code]: "acyclic r \ irrefl (r\<^sup>+)" by (simp add: acyclic_def irrefl_def) lemma acyclicI: "\x. (x, x) \ r\<^sup>+ \ acyclic r" by (simp add: acyclic_def) lemma (in preorder) acyclicI_order: assumes *: "\a b. (a, b) \ r \ f b < f a" shows "acyclic r" proof - have "f b < f a" if "(a, b) \ r\<^sup>+" for a b using that by induct (auto intro: * less_trans) then show ?thesis by (auto intro!: acyclicI) qed lemma acyclic_insert [iff]: "acyclic (insert (y, x) r) \ acyclic r \ (x, y) \ r\<^sup>*" by (simp add: acyclic_def trancl_insert) (blast intro: rtrancl_trans) lemma acyclic_converse [iff]: "acyclic (r\) \ acyclic r" by (simp add: acyclic_def trancl_converse) lemmas acyclicP_converse [iff] = acyclic_converse [to_pred] lemma acyclic_impl_antisym_rtrancl: "acyclic r \ antisym (r\<^sup>*)" by (simp add: acyclic_def antisym_def) (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl) (* Other direction: acyclic = no loops antisym = only self loops Goalw [acyclic_def,antisym_def] "antisym( r\<^sup>* ) \ acyclic(r - Id) \ antisym( r\<^sup>* ) = acyclic(r - Id)"; *) lemma acyclic_subset: "acyclic s \ r \ s \ acyclic r" unfolding acyclic_def by (blast intro: trancl_mono) subsection \Setup of transitivity reasoner\ ML \ structure Trancl_Tac = Trancl_Tac ( val r_into_trancl = @{thm trancl.r_into_trancl}; val trancl_trans = @{thm trancl_trans}; val rtrancl_refl = @{thm rtrancl.rtrancl_refl}; val r_into_rtrancl = @{thm r_into_rtrancl}; val trancl_into_rtrancl = @{thm trancl_into_rtrancl}; val rtrancl_trancl_trancl = @{thm rtrancl_trancl_trancl}; val trancl_rtrancl_trancl = @{thm trancl_rtrancl_trancl}; val rtrancl_trans = @{thm rtrancl_trans}; fun decomp (\<^const>\Trueprop\ $ t) = let fun dec (Const (\<^const_name>\Set.member\, _) $ (Const (\<^const_name>\Pair\, _) $ a $ b) $ rel) = let fun decr (Const (\<^const_name>\rtrancl\, _ ) $ r) = (r,"r*") | decr (Const (\<^const_name>\trancl\, _ ) $ r) = (r,"r+") | decr r = (r,"r"); val (rel,r) = decr (Envir.beta_eta_contract rel); in SOME (a,b,rel,r) end | dec _ = NONE in dec t end | decomp _ = NONE; ); structure Tranclp_Tac = Trancl_Tac ( val r_into_trancl = @{thm tranclp.r_into_trancl}; val trancl_trans = @{thm tranclp_trans}; val rtrancl_refl = @{thm rtranclp.rtrancl_refl}; val r_into_rtrancl = @{thm r_into_rtranclp}; val trancl_into_rtrancl = @{thm tranclp_into_rtranclp}; val rtrancl_trancl_trancl = @{thm rtranclp_tranclp_tranclp}; val trancl_rtrancl_trancl = @{thm tranclp_rtranclp_tranclp}; val rtrancl_trans = @{thm rtranclp_trans}; fun decomp (\<^const>\Trueprop\ $ t) = let fun dec (rel $ a $ b) = let fun decr (Const (\<^const_name>\rtranclp\, _ ) $ r) = (r,"r*") | decr (Const (\<^const_name>\tranclp\, _ ) $ r) = (r,"r+") | decr r = (r,"r"); val (rel,r) = decr rel; in SOME (a, b, rel, r) end | dec _ = NONE in dec t end | decomp _ = NONE; ); \ setup \ map_theory_simpset (fn ctxt => ctxt addSolver (mk_solver "Trancl" Trancl_Tac.trancl_tac) addSolver (mk_solver "Rtrancl" Trancl_Tac.rtrancl_tac) addSolver (mk_solver "Tranclp" Tranclp_Tac.trancl_tac) addSolver (mk_solver "Rtranclp" Tranclp_Tac.rtrancl_tac)) \ +lemma transp_rtranclp [simp]: "transp R\<^sup>*\<^sup>*" + by(auto simp add: transp_def) text \Optional methods.\ method_setup trancl = \Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.trancl_tac)\ \simple transitivity reasoner\ method_setup rtrancl = \Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.rtrancl_tac)\ \simple transitivity reasoner\ method_setup tranclp = \Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.trancl_tac)\ \simple transitivity reasoner (predicate version)\ method_setup rtranclp = \Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.rtrancl_tac)\ \simple transitivity reasoner (predicate version)\ end